]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/freescale/bool_lemmas.ma
freescale porting to ng, work in progress
[helm.git] / helm / software / matita / contribs / ng_assembly / freescale / bool_lemmas.ma
index 5d05eb19575e34ee6197ca328355e9b14de26402..b03e5a62ddc181cc0378befe0728cf3cebd2f451 100755 (executable)
@@ -31,12 +31,14 @@ include "freescale/bool.ma".
 (* BOOLEANI *)
 (* ******** *)
 
-ndefinition bool_destruct :
- Πb1,b2:bool.ΠP:Prop.b1 = b2 →
+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 ]
   ].
+
+ndefinition bool_destruct : bool_destruct_aux.
  #b1; #b2; #P;
  nelim b1;
  nelim b2;
@@ -116,67 +118,61 @@ nlemma associative_orbool : ∀b1,b2,b3.((b1 ⊙ b2) ⊙ b3) = (b1 ⊙ (b2 ⊙ b
 nqed.
 
 nlemma eqbool_to_eq : ∀b1,b2:bool.(eq_bool b1 b2 = true) → (b1 = b2).
- #b1; #b2; #H;
- nletin K ≝ (bool_destruct ?? (b1 = b2) H);
- nelim b1 in K:(%) ⊢ %;
- nelim b2;
+ #b1; #b2;
+ ncases b1;
+ ncases b2;
  nnormalize;
- ##[ ##2,3: #H; napply H
- ##| ##1,4: #H; napply (refl_eq ??)
+ ##[ ##1,4: #H; napply (refl_eq ??)
+ ##| ##*: #H; napply (bool_destruct ??? H)
  ##]
 nqed.
 
 nlemma eq_to_eqbool : ∀b1,b2.b1 = b2 → eq_bool b1 b2 = true.
- #b1; #b2; #H;
- nletin K ≝ (bool_destruct ?? (eq_bool b1 b2 = true) H);
- nelim b1 in K:(%) ⊢ %;
- nelim b2;
+ #b1; #b2;
+ ncases b1;
+ ncases b2;
  nnormalize;
- ##[ ##2,3: #H; napply H
- ##| ##1,4: #H; napply (refl_eq ??)
+ ##[ ##1,4: #H; napply (refl_eq ??)
+ ##| ##*: #H; napply (bool_destruct ??? H)
  ##]
 nqed.
 
 nlemma andb_true_true_l: ∀b1,b2.(b1 ⊗ b2) = true → b1 = true.
- #b1; #b2; #H;
- nletin K ≝ (bool_destruct ?? (b1 = true) H);
- nelim b1 in K:(%) ⊢ %;
- nelim b2;
+ #b1; #b2;
+ ncases b1;
+ ncases b2;
  nnormalize;
- ##[ ##3,4: #H; napply H
- ##| ##1,2: #H; napply (refl_eq ??)
+ ##[ ##1,2: #H; napply (refl_eq ??)
+ ##| ##*: #H; napply (bool_destruct ??? H)
  ##]
 nqed.
 
 nlemma andb_true_true_r: ∀b1,b2.(b1 ⊗ b2) = true → b2 = true.
- #b1; #b2; #H;
- nletin K ≝ (bool_destruct ?? (b2 = true) H);
- nelim b1 in K:(%) ⊢ %;
- nelim b2;
+ #b1; #b2;
+ ncases b1;
+ ncases b2;
  nnormalize;
- ##[ ##2,4: #H; napply H
- ##| ##1,3: #H; napply (refl_eq ??)
+ ##[ ##1,3: #H; napply (refl_eq ??)
+ ##| ##*: #H; napply (bool_destruct ??? H)
  ##]
 nqed.
 
 nlemma orb_false_false_l : ∀b1,b2:bool.(b1 ⊕ b2) = false → b1 = false.
- #b1; #b2; #H;
- nletin K ≝ (bool_destruct ?? (b1 = false) H);
- nelim b1 in K:(%) ⊢ %;
- nelim b2;
+ #b1; #b2;
+ ncases b1;
+ ncases b2;
  nnormalize;
- ##[ ##1,2,3: #H; napply H
- ##| ##4: #H; napply (refl_eq ??)
+ ##[ ##4: #H; napply (refl_eq ??)
+ ##| ##*: #H; napply (bool_destruct ??? H)
  ##]
 nqed.
 
 nlemma orb_false_false_r : ∀b1,b2:bool.(b1 ⊕ b2) = false → b2 = false.
- #b1; #b2; #H;
- nletin K ≝ (bool_destruct ?? (b2 = false) H);
- nelim b1 in K:(%) ⊢ %;
- nelim b2;
+ #b1; #b2;
+ ncases b1;
+ ncases b2;
  nnormalize;
- ##[ ##1,2,3: #H; napply H
- ##| ##4: #H; napply (refl_eq ??)
+ ##[ ##4: #H; napply (refl_eq ??)
+ ##| ##*: #H; napply (bool_destruct ??? H)
  ##]
 nqed.