X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fassembly%2Ffreescale%2Fextra.ma;h=207a5a267d386b67af511cedf815e1c242114261;hb=026c6c5b0e094b2e6e8244909bc5ac3d88b70b9c;hp=bc76c4d6b7f0addf954b61fa572c3ad3657d803e;hpb=9eabe046c1182960de8cfdba96c5414224e3a61e;p=helm.git diff --git a/helm/software/matita/contribs/assembly/freescale/extra.ma b/helm/software/matita/contribs/assembly/freescale/extra.ma index bc76c4d6b..207a5a267 100644 --- a/helm/software/matita/contribs/assembly/freescale/extra.ma +++ b/helm/software/matita/contribs/assembly/freescale/extra.ma @@ -54,6 +54,71 @@ definition eq_bool ≝ [ true ⇒ b2 | false ⇒ not_bool b2 ]. +lemma eqbool_switch : ∀b1,b2.eq_bool b1 b2 = eq_bool b2 b1. + do 2 intro; + elim b1; elim b2; + reflexivity. +qed. + +lemma andbool_switch : ∀b1,b2.and_bool b1 b2 = and_bool b2 b1. + do 2 intro; + elim b1; elim b2; + reflexivity. +qed. + +lemma orbool_switch : ∀b1,b2.or_bool b1 b2 = or_bool b2 b1. + do 2 intro; + elim b1; elim b2; + reflexivity. +qed. + +lemma xorbool_switch : ∀b1,b2.xor_bool b1 b2 = xor_bool b2 b1. + do 2 intro; + elim b1; elim b2; + reflexivity. +qed. + + +lemma orb_false_false : + ∀b1,b2:bool.((or_bool b1 b2) = false) → b1 = false. + intros 2; + elim b1 0; + elim b2; + simplify in H; + try destruct H; + reflexivity. +qed. + +lemma orb_false_false_r : + ∀b1,b2:bool.((or_bool b1 b2) = false) → b2 = false. + intros 2; + elim b1 0; + elim b2; + simplify in H; + try destruct H; + reflexivity. +qed. + +lemma eqbool_to_eq : ∀b1,b2:bool.(eq_bool b1 b2 = true) → (b1 = b2). + unfold eq_bool; + intros; + elim b1 in H:(%); + elim b2 in H:(%); + normalize in H:(%); + try reflexivity; + destruct H. +qed. + +lemma eq_to_eqbool : ∀b1,b2.b1 = b2 → eq_bool b1 b2 = true. + do 2 intro; + elim b1 0; + elim b2 0; + intro; + normalize in H:(%); + try destruct H; + reflexivity. +qed. + (* \ominus *) notation "hvbox(⊖ a)" non associative with precedence 36 for @{ 'not_bool $a }.