X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Flists%2Flistb.ma;h=a95d8a05adb91bcb3eeb7054f200d4759cfe632c;hb=3447747453bbf439d071d21dcb68149cae3a9068;hp=ac72b008bee54237a9360bb2830545ab5ec7b853;hpb=347a03a55e60c5a4682e0133454ff87ce21a4e8c;p=helm.git diff --git a/matita/matita/lib/basics/lists/listb.ma b/matita/matita/lib/basics/lists/listb.ma index ac72b008b..a95d8a05a 100644 --- a/matita/matita/lib/basics/lists/listb.ma +++ b/matita/matita/lib/basics/lists/listb.ma @@ -103,6 +103,16 @@ 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 @@ -112,6 +122,15 @@ lemma mem_to_memb: ∀S:DeqSet.∀a,l. mem S a l → memb S a l = true. ] ] qed. + +lemma memb_to_mem: ∀S:DeqSet.∀l,a. memb S a l =true → mem S a l. +#S #l #a elim l + [normalize #H destruct + |#b #tl #Hind #mema cases (orb_true_l … mema) + [#eqab >(\P eqab) %1 % |#memtl %2 @Hind @memtl] + ] +qed. + (**************** unicity test *****************) let rec uniqueb (S:DeqSet) l on l : bool ≝