]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/basics/bool.ma
backport of WIP on \lambda\delta to matita 0.99.3
[helm.git] / matita / matita / lib / basics / bool.ma
index 43a3369a393d0c9e4725610beac64ceb868b6e6a..0d790efc54c590f0e5e8b4f47257d5b9ad99ad3f 100644 (file)
@@ -37,6 +37,14 @@ theorem notb_notb: ∀b:bool. notb (notb b) = b.
 theorem injective_notb: injective bool bool notb.
 #b1 #b2 #H // qed.
 
+theorem noteq_to_eqnot: ∀b1,b2. b1 ≠ b2 → b1 = notb b2.
+* * // #H @False_ind /2/
+qed.
+
+theorem eqnot_to_noteq: ∀b1,b2. b1 = notb b2 → b1 ≠ b2.
+* * normalize // #_ @(not_to_not … not_eq_true_false) //
+qed.
+
 definition andb : bool → bool → bool ≝
 λb1,b2:bool. match b1 with [ true ⇒ b2 | false ⇒ false ].
 
@@ -46,6 +54,10 @@ theorem andb_elim: ∀ b1,b2:bool. ∀ P:bool → Prop.
 match b1 with [ true ⇒ P b2 | false ⇒ P false] → P (b1 ∧ b2).
 #b1 #b2 #P (elim b1) normalize // qed.
 
+theorem true_to_andb_true: ∀b1,b2. b1 = true → b2 = true → (b1 ∧ b2) = true.
+#b1 cases b1 normalize //
+qed.
+
 theorem andb_true_l: ∀ b1,b2. (b1 ∧ b2) = true → b1 = true.
 #b1 (cases b1) normalize // qed.
 
@@ -85,7 +97,7 @@ definition xorb : bool → bool → bool ≝
 notation > "'if' term 46 e 'then' term 46 t 'else' term 46 f" non associative with precedence 46
  for @{ match $e in bool with [ true ⇒ $t | false ⇒ $f]  }.
 notation < "hvbox('if' \nbsp term 46 e \nbsp break 'then' \nbsp term 46 t \nbsp break 'else' \nbsp term 49 f \nbsp)" non associative with precedence 46
- for @{ match $e with [ true ⇒ $t | false ⇒ $f]  }.
+ for @{ match $e return $T with [ true ⇒ $t | false ⇒ $f]  }.
 
 theorem bool_to_decidable_eq: 
   ∀b1,b2:bool. decidable (b1=b2).
@@ -95,3 +107,4 @@ theorem true_or_false:
 ∀b:bool. b = true ∨ b = false.
 #b (cases b) /2/ qed.
 
+