]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/tutorial/chapter6.ma
decentralizing core notation
[helm.git] / matita / matita / lib / tutorial / chapter6.ma
index 258152d43d204a3a62806027381f31e2047f052d..a5355b5fa990fdea2185197ca25a45553ff5e5d9 100644 (file)
@@ -1,5 +1,6 @@
 
 include "tutorial/chapter5.ma".
+include "basics/core_notation/singl_1.ma".
 
 (*************************** Naive Set Theory *********************************)
 
@@ -204,6 +205,20 @@ notation "\P H" non associative with precedence 90
 notation "\b H" non associative with precedence 90 
   for @{(proj2 … (eqb_true ???) $H)}. 
   
+lemma eqb_false: ∀S:DeqSet.∀a,b:S. 
+  (eqb ? a b) = false ↔ a ≠ b.
+#S #a #b % #H 
+  [@(not_to_not … not_eq_true_false) #H1 <H @sym_eq @(\b H1)
+  |cases (true_or_false (eqb ? a b)) // #H1 @False_ind @(absurd … (\P H1) H)
+  ]
+qed.
+notation "\Pf H" non associative with precedence 90 
+  for @{(proj1 … (eqb_false ???) $H)}. 
+
+notation "\bf H" non associative with precedence 90 
+  for @{(proj2 … (eqb_false ???) $H)}. 
+
 (****************************** Unification hints *****************************)
 
 (* Suppose now to write an expression of the following kind:
@@ -329,6 +344,20 @@ let rec memb (S:DeqSet) (x:S) (l: list S) on l  ≝
   | cons a tl ⇒ (x == a) ∨ memb S x tl
   ].
   
+lemma memb_hd: ∀S,a,l. memb S a (a::l) = true.
+#S #a #l normalize >(proj2 … (eqb_true S …) (refl S a)) //
+qed.
+
+lemma memb_cons: ∀S,a,b,l. 
+  memb S a l = true → memb S a (b::l) = true.
+#S #a #b #l normalize cases (a==b) normalize // 
+qed.
+
+lemma memb_single: ∀S,a,x. memb S a [x] = true → a = x.
+#S #a #x normalize cases (true_or_false … (a==x)) #H
+  [#_ >(\P H) // |>H normalize #abs @False_ind /2/]
+qed.
+
 let rec uniqueb (S:DeqSet) l on l : bool ≝
   match l with 
   [ nil ⇒ true
@@ -391,7 +420,7 @@ lemma memb_map_to_exists: ∀A,B:DeqSet.∀f:A→B.∀l,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) % // normalize >(\b (refl ? a)) //
+    [#eqb @(ex_intro … a) <(\P eqb) % // 
     |#memb cases (Hind … memb) #a * #mema #eqb
      @(ex_intro … a) /3/
     ]