]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/basics/lists/listb.ma
A few integrations (closed an axiom in finset).
[helm.git] / matita / matita / lib / basics / lists / listb.ma
index e46f45f075416af0fd6b58c18fe2695e2da61515..4413e6ed1a8e68bec27302f1938b3156b7ffc998 100644 (file)
@@ -31,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)) //
@@ -104,6 +103,34 @@ 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.
+
+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 ≝
@@ -145,6 +172,32 @@ cases (true_or_false … (memb S a (unique_append S tl l2)))
 #H >H normalize [@Hind //] >H normalize @Hind //
 qed.
 
+lemma uniqueb_append: ∀A,l1,l2. uniqueb A l1 = true → uniqueb A l2 =true → 
+  (∀a. memb A a l1 =true → ¬ memb A a l2 =true) → uniqueb A (l1@l2) = true.
+#A #l1 elim l1 [normalize //] #a #tl #Hind #l2 #Hatl #Hl2 
+#Hmem normalize cut (memb A a (tl@l2)=false)
+  [2:#Hcut >Hcut normalize @Hind //
+    [@(andb_true_r … Hatl) |#x #Hmemx @Hmem @orb_true_r2 //]
+  |@(noteq_to_eqnot ? true) % #Happend cases (memb_append … Happend)
+    [#H1 @(absurd … H1) @sym_not_eq @eqnot_to_noteq 
+     @sym_eq @(andb_true_l … Hatl)
+    |#H @(absurd … H) @Hmem normalize >(\b (refl ? a)) //
+    ]
+  ]
+qed.
+
+lemma memb_map_to_exists: ∀A,B:DeqSet.∀f:A→B.∀l,b. 
+  memb ? b (map ?? f l) = true → ∃a. memb ? a l = true ∧ f a = b.
+#A #B #f #l elim l 
+  [#b normalize #H destruct (H) 
+  |#a #tl #Hind #b #H cases (orb_true_l … H) 
+    [#eqb @(ex_intro … a) <(\P eqb) % // 
+    |#memb cases (Hind … memb) #a * #mema #eqb
+     @(ex_intro … a) /3/
+    ]
+  ]
+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
@@ -226,6 +279,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) <H // 
+    |/2/
+    ]
+  |>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/]
@@ -235,13 +309,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.