]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/num/bool_lemmas.ma
mod change (-x)
[helm.git] / helm / software / matita / contribs / ng_assembly / num / bool_lemmas.ma
old mode 100755 (executable)
new mode 100644 (file)
index 9190add..db5214c
@@ -16,7 +16,7 @@
 (*                          Progetto FreeScale                            *)
 (*                                                                        *)
 (*   Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it              *)
-(*   Ultima modifica: 05/08/2009                                          *)
+(*   Sviluppo: 2008-2010                                                  *)
 (*                                                                        *)
 (* ********************************************************************** *)
 
@@ -26,32 +26,19 @@ include "num/bool.ma".
 (* BOOLEANI *)
 (* ******** *)
 
+(*
 ndefinition bool_destruct_aux ≝
 Πb1,b2:bool.ΠP:Prop.b1 = b2 →
- match b1 with
-  [ true ⇒ match b2 with [ true ⇒ P → P | false ⇒ P ]
-  | false ⇒ match b2 with [ true ⇒ P | false ⇒ P → P ]
-  ].
+ match eq_bool b1 b2 with [ true ⇒ P → P | false ⇒ P ].
 
 ndefinition bool_destruct : bool_destruct_aux.
- #b1; #b2; #P;
+ #b1; #b2; #P; #H;
+ nrewrite < H;
  nelim b1;
- nelim b2;
  nnormalize;
- #H;
- ##[ ##2: napply False_ind;
-          nchange with (match true with [ true ⇒ False | false ⇒ True]);
-          nrewrite > H;
-          nnormalize;
-          napply I
- ##| ##3: napply False_ind;
-          nchange with (match true with [ true ⇒ False | false ⇒ True]);
-          nrewrite < H;
-          nnormalize;
-          napply I
- ##| ##1,4: napply (λx:P.x)
- ##]
+ napply (λx.x).
 nqed.
+*)
 
 nlemma symmetric_eqbool : symmetricT bool bool eq_bool.
  #b1; #b2;
@@ -118,7 +105,7 @@ nlemma eqbool_to_eq : ∀b1,b2:bool.(eq_bool b1 b2 = true) → (b1 = b2).
  ncases b2;
  nnormalize;
  ##[ ##1,4: #H; napply refl_eq
- ##| ##*: #H; napply (bool_destruct … H)
+ ##| ##*: #H; ndestruct (*napply (bool_destruct … H)*)
  ##]
 nqed.
 
@@ -128,7 +115,7 @@ nlemma eq_to_eqbool : ∀b1,b2.b1 = b2 → eq_bool b1 b2 = true.
  ncases b2;
  nnormalize;
  ##[ ##1,4: #H; napply refl_eq
- ##| ##*: #H; napply (bool_destruct … H)
+ ##| ##*: #H; ndestruct (*napply (bool_destruct … H)*)
  ##]
 nqed.
 
@@ -138,7 +125,17 @@ nlemma decidable_bool : ∀x,y:bool.decidable (x = y).
  nelim x;
  nelim y;
  ##[ ##1,4: napply (or2_intro1 (? = ?) (? ≠ ?) …); napply refl_eq
- ##| ##*: napply (or2_intro2 (? = ?) (? ≠ ?) …); nnormalize; #H; napply False_ind; napply (bool_destruct … H)
+ ##| ##*: napply (or2_intro2 (? = ?) (? ≠ ?) …);
+          nnormalize; #H;
+          napply False_ind;
+          ndestruct (*napply (bool_destruct … H)*)
+ ##]
+nqed.
+
+nlemma decidable_bexpr : ∀x.(x = true) ∨ (x = false).
+ #x; ncases x;
+ ##[ ##1: napply (or2_intro1 (true = true) (true = false) (refl_eq …))
+ ##| ##2: napply (or2_intro2 (false = true) (false = false) (refl_eq …))
  ##]
 nqed.
 
@@ -147,8 +144,8 @@ nlemma neqbool_to_neq : ∀b1,b2:bool.(eq_bool b1 b2 = false) → (b1 ≠ b2).
  ncases b1;
  ncases b2;
  nnormalize;
- ##[ ##1,4: #H; napply (bool_destruct … H)
- ##| ##*: #H; #H1; napply (bool_destruct … H1)
+ ##[ ##1,4: #H; ndestruct (*napply (bool_destruct … H)*)
+ ##| ##*: #H; #H1; ndestruct (*napply (bool_destruct … H1)*)
  ##]
 nqed.
 
@@ -162,13 +159,41 @@ nlemma neq_to_neqbool : ∀b1,b2.b1 ≠ b2 → eq_bool b1 b2 = false.
  ##]
 nqed.
 
+nlemma eqfalse_to_neqtrue : ∀x.x = false → x ≠ true.
+ #x; nelim x;
+ ##[ ##1: #H; ndestruct (*napply (bool_destruct … H)*)
+ ##| ##2: #H; nnormalize; #H1; ndestruct (*napply (bool_destruct … H1)*)
+ ##]
+nqed.
+
+nlemma eqtrue_to_neqfalse : ∀x.x = true → x ≠ false.
+ #x; nelim x;
+ ##[ ##1: #H; nnormalize; #H1; ndestruct (*napply (bool_destruct … H1)*)
+ ##| ##2: #H; ndestruct (*napply (bool_destruct … H)*)
+ ##]
+nqed.
+
+nlemma neqfalse_to_eqtrue : ∀x.x ≠ false → x = true.
+ #x; nelim x;
+ ##[ ##1: #H; napply refl_eq
+ ##| ##2: nnormalize; #H; nelim (H (refl_eq …))
+ ##]
+nqed.
+
+nlemma neqtrue_to_eqfalse : ∀x.x ≠ true → x = false.
+ #x; nelim x;
+ ##[ ##1: nnormalize; #H; nelim (H (refl_eq …))
+ ##| ##2: #H; napply refl_eq
+ ##]
+nqed.
+
 nlemma andb_true_true_l: ∀b1,b2.(b1 ⊗ b2) = true → b1 = true.
  #b1; #b2;
  ncases b1;
  ncases b2;
  nnormalize;
  ##[ ##1,2: #H; napply refl_eq
- ##| ##*: #H; napply (bool_destruct … H)
+ ##| ##*: #H; ndestruct (*napply (bool_destruct … H)*)
  ##]
 nqed.
 
@@ -178,7 +203,7 @@ nlemma andb_true_true_r: ∀b1,b2.(b1 ⊗ b2) = true → b2 = true.
  ncases b2;
  nnormalize;
  ##[ ##1,3: #H; napply refl_eq
- ##| ##*: #H; napply (bool_destruct … H)
+ ##| ##*: #H; ndestruct (*napply (bool_destruct … H)*)
  ##]
 nqed.
 
@@ -189,7 +214,7 @@ nlemma andb_false2
  ncases b1;
  ncases b2;
  nnormalize;
- ##[ ##1: #H; napply (bool_destruct … H)
+ ##[ ##1: #H; ndestruct (*napply (bool_destruct … H)*)
  ##| ##2,4: #H; napply (or2_intro2 … H)
  ##| ##3: #H; napply (or2_intro1 … H)
  ##]
@@ -203,7 +228,7 @@ nlemma andb_false3
  ncases b2;
  ncases b3;
  nnormalize;
- ##[ ##1: #H; napply (bool_destruct … H)
+ ##[ ##1: #H; ndestruct (*napply (bool_destruct … H)*)
  ##| ##5,6,7,8: #H; napply (or3_intro1 … H)
  ##| ##2,4: #H; napply (or3_intro3 … H)
  ##| ##3: #H; napply (or3_intro2 … H)
@@ -219,7 +244,7 @@ nlemma andb_false4
  ncases b3;
  ncases b4;
  nnormalize;
- ##[ ##1: #H; napply (bool_destruct … H)
+ ##[ ##1: #H; ndestruct (*napply (bool_destruct … H)*)
  ##| ##9,10,11,12,13,14,15,16: #H; napply (or4_intro1 … H)
  ##| ##5,6,7,8: #H; napply (or4_intro2 … H)
  ##| ##3,4: #H; napply (or4_intro3 … H)
@@ -237,7 +262,7 @@ nlemma andb_false5
  ncases b4;
  ncases b5;
  nnormalize;
- ##[ ##1: #H; napply (bool_destruct … H)
+ ##[ ##1: #H; ndestruct (*napply (bool_destruct … H)*)
  ##| ##17,18,19,20,21,22,23,24,25,26,27,28,29,30,31,32: #H; napply (or5_intro1 … H)
  ##| ##9,10,11,12,13,14,15,16: #H; napply (or5_intro2 … H)
  ##| ##5,6,7,8: #H; napply (or5_intro3 … H)
@@ -284,7 +309,7 @@ nlemma orb_false_false_l : ∀b1,b2:bool.(b1 ⊕ b2) = false → b1 = false.
  ncases b2;
  nnormalize;
  ##[ ##4: #H; napply refl_eq
- ##| ##*: #H; napply (bool_destruct … H)
+ ##| ##*: #H; ndestruct (*napply (bool_destruct … H)*)
  ##]
 nqed.
 
@@ -294,6 +319,6 @@ nlemma orb_false_false_r : ∀b1,b2:bool.(b1 ⊕ b2) = false → b2 = false.
  ncases b2;
  nnormalize;
  ##[ ##4: #H; napply refl_eq
- ##| ##*: #H; napply (bool_destruct … H)
+ ##| ##*: #H; ndestruct (*napply (bool_destruct … H)*)
  ##]
 nqed.