X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Ffreescale%2Fbool_lemmas.ma;h=493a8e5b6524af78bf0abbfd62219340b23ac35a;hb=c3fc204cc02f1031b5d17fb0f2be1fc01e5c452f;hp=5d05eb19575e34ee6197ca328355e9b14de26402;hpb=3e5c359c75874748cfed8a9046031b62396e0e6d;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/freescale/bool_lemmas.ma b/helm/software/matita/contribs/ng_assembly/freescale/bool_lemmas.ma index 5d05eb195..493a8e5b6 100755 --- a/helm/software/matita/contribs/ng_assembly/freescale/bool_lemmas.ma +++ b/helm/software/matita/contribs/ng_assembly/freescale/bool_lemmas.ma @@ -116,67 +116,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.