X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Flists%2Flistb.ma;h=438f3b497f5ed5bbc74789af9bb81d555a10e0d8;hb=b0d97cd7e2c50fb1fc2d50c86f3140e226b08a81;hp=5abf50f14274a39f3330ff7da75995a8bce0f8e4;hpb=734814ac0ecee6e2e6db6637bbd6acb88b505585;p=helm.git diff --git a/matita/matita/lib/basics/lists/listb.ma b/matita/matita/lib/basics/lists/listb.ma index 5abf50f14..438f3b497 100644 --- a/matita/matita/lib/basics/lists/listb.ma +++ b/matita/matita/lib/basics/lists/listb.ma @@ -16,6 +16,13 @@ include "basics/lists/list.ma". include "basics/sets.ma". include "basics/deqsets.ma". +(********* isnilb *********) +let rec isnilb A (l: list A) on l ≝ +match l with +[ nil ⇒ true +| cons hd tl ⇒ false +]. + (********* search *********) let rec memb (S:DeqSet) (x:S) (l: list S) on l ≝ @@ -24,8 +31,7 @@ let rec memb (S:DeqSet) (x:S) (l: list S) on l ≝ | cons a tl ⇒ (x == a) ∨ memb S x tl ]. -notation < "\memb x l" non associative with precedence 90 for @{'memb $x $l}. -interpretation "boolean membership" 'memb a l = (memb ? a l). +interpretation "boolean membership" 'mem a l = (memb ? a l). lemma memb_hd: ∀S,a,l. memb S a (a::l) = true. #S #a #l normalize >(proj2 … (eqb_true S …) (refl S a)) // @@ -97,6 +103,25 @@ lemma memb_compose: ∀S1,S2,S3,op,a1,a2,l1,l2. ] qed. +lemma memb_reverse: ∀S:DeqSet.∀a:S.∀l. + memb ? a l = true → memb ? a (reverse ? l) = true. +#S #a #l elim l [normalize //] +#b #tl #Hind #memba change with ([b]@tl) in match (b::tl); +>reverse_append cases (orb_true_l … memba) #Hcase + [@memb_append_l2 >(\P Hcase) whd in match (reverse ??); @memb_hd + |@memb_append_l1 /2/ + ] +qed. + +lemma mem_to_memb: ∀S:DeqSet.∀a,l. mem S a l → memb S a l = true. +#S #a #l elim l normalize + [@False_ind + |#hd #tl #Hind * + [#eqa >(\b eqa) % + |#Hmem cases (a==hd) // normalize /2/ + ] + ] +qed. (**************** unicity test *****************) let rec uniqueb (S:DeqSet) l on l : bool ≝ @@ -138,6 +163,30 @@ cases (true_or_false … (memb S a (unique_append S tl l2))) #H >H normalize [@Hind //] >H normalize @Hind // qed. +lemma memb_map_inj: ∀A,B:DeqSet.∀f:A→B.∀l,a. injective A B f → + memb ? (f a) (map ?? f l) = true → memb ? a l = true. +#A #B #f #l #a #injf elim l + [normalize // + |#a1 #tl #Hind #Ha cases (orb_true_l … Ha) + [#eqf @orb_true_r1 @(\b ?) >(injf … (\P eqf)) // + |#membtl @orb_true_r2 /2/ + ] + ] +qed. + +lemma unique_map_inj: ∀A,B:DeqSet.∀f:A→B.∀l. injective A B f → + uniqueb A l = true → uniqueb B (map ?? f l) = true . +#A #B #f #l #injf elim l + [normalize // + |#a #tl #Hind #Htl @true_to_andb_true + [@sym_eq @noteq_to_eqnot @sym_not_eq + @(not_to_not ?? (memb_map_inj … injf …) ) + <(andb_true_l ?? Htl) /2/ + |@Hind @(andb_true_r ?? Htl) + ] + ] +qed. + (******************* sublist *******************) definition sublist ≝ λS,l1,l2.∀a. memb S a l1 = true → memb S a l2 = true. @@ -195,6 +244,27 @@ qed. (********************* filtering *****************) +lemma memb_filter_memb: ∀S,f,a,l. + memb S a (filter S f l) = true → memb S a l = true. +#S #f #a #l elim l [normalize //] +#b #tl #Hind normalize (cases (f b)) normalize +cases (a==b) normalize // @Hind +qed. + +lemma uniqueb_filter : ∀S,l,f. + uniqueb S l = true → uniqueb S (filter S f l) = true. +#S #l #f elim l // +#a #tl #Hind #Hunique cases (andb_true … Hunique) +#memba #uniquetl cases (true_or_false … (f a)) #Hfa + [>filter_true // @true_to_andb_true + [@sym_eq @noteq_to_eqnot @(not_to_not … (eqnot_to_noteq … (sym_eq … memba))) + #H @sym_eq @(memb_filter_memb … f) filter_false /2/ + ] +qed. + lemma filter_true: ∀S,f,a,l. memb S a (filter S f l) = true → f a = true. #S #f #a #l elim l [normalize #H @False_ind /2/] @@ -204,13 +274,6 @@ cases (true_or_false (a==b)) #eqab [#_ >(\P eqab) // | >eqab normalize @Hind] qed. -lemma memb_filter_memb: ∀S,f,a,l. - memb S a (filter S f l) = true → memb S a l = true. -#S #f #a #l elim l [normalize //] -#b #tl #Hind normalize (cases (f b)) normalize -cases (a==b) normalize // @Hind -qed. - lemma memb_filter: ∀S,f,l,x. memb S x (filter ? f l) = true → memb S x l = true ∧ (f x = true). /3/ qed.