X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fassembly%2Ffreescale%2Fextra.ma;h=7da45c848e3a2aef8e3f012548998ea217c177ce;hb=db020b4218272e2e35641ce3bc3b0a9b3afda899;hp=f683ecf8fa42e271a7a252c04a3d0ce507ffcb7c;hpb=55e5bef77f163b29feeb9ad4e83376c5aa301297;p=helm.git diff --git a/helm/software/matita/contribs/assembly/freescale/extra.ma b/helm/software/matita/contribs/assembly/freescale/extra.ma index f683ecf8f..7da45c848 100644 --- a/helm/software/matita/contribs/assembly/freescale/extra.ma +++ b/helm/software/matita/contribs/assembly/freescale/extra.ma @@ -78,6 +78,27 @@ lemma xorbool_switch : ∀b1,b2.xor_bool b1 b2 = xor_bool b2 b1. 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; @@ -187,7 +208,7 @@ axiom or_lt_le : ∀n,m. n < m ∨ m ≤ n. lemma le_to_lt: ∀n,m. n ≤ m → n < S m. intros; - autobatch. + unfold;autobatch. qed. alias num (instance 0) = "natural number". @@ -232,7 +253,7 @@ lemma leq_m_n_to_eq_div_n_m_S: ∀n,m:nat. 0 < m → m ≤ n → ∃z. n/m = S z ] | elim H1; autobatch ] - | autobatch + | exists;[apply (pred m);]autobatch ]. qed.