]> matita.cs.unibo.it Git - helm.git/commitdiff
fixing
authorAndrea Asperti <andrea.asperti@unibo.it>
Thu, 4 Dec 2014 09:36:22 +0000 (09:36 +0000)
committerAndrea Asperti <andrea.asperti@unibo.it>
Thu, 4 Dec 2014 09:36:22 +0000 (09:36 +0000)
matita/matita/lib/tutorial/chapter6.ma

index 258152d43d204a3a62806027381f31e2047f052d..fb38f2d316e8d9db3ec58a345f34ee2ae513dd5a 100644 (file)
@@ -204,6 +204,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: