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=7951f2997265484adf6ec17130e2148f9f043420;hpb=1f946a70bc439ced80c695f0fd0c210df0d3b767;p=helm.git diff --git a/matita/matita/lib/basics/lists/listb.ma b/matita/matita/lib/basics/lists/listb.ma index 7951f2997..438f3b497 100644 --- a/matita/matita/lib/basics/lists/listb.ma +++ b/matita/matita/lib/basics/lists/listb.ma @@ -103,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 ≝ @@ -225,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/] @@ -234,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.