]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/assembly/freescale/extra.ma
freescale porting
[helm.git] / helm / software / matita / contribs / assembly / freescale / extra.ma
index 94f4a47547eebf1a1048c94474c14f35e4e90625..7da45c848e3a2aef8e3f012548998ea217c177ce 100644 (file)
@@ -54,29 +54,90 @@ 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 }.
-interpretation "not_bool" 'not_bool x = 
- (cic:/matita/freescale/extra/not_bool.con x).
+interpretation "not_bool" 'not_bool x = (not_bool x).
 
 (* \otimes *)
 notation "hvbox(a break ⊗ b)" left associative with precedence 35
  for @{ 'and_bool $a $b }.
-interpretation "and_bool" 'and_bool x y = 
- (cic:/matita/freescale/extra/and_bool.con x y).
+interpretation "and_bool" 'and_bool x y = (and_bool x y).
 
 (* \oplus *)
 notation "hvbox(a break ⊕ b)" left associative with precedence 34
  for @{ 'or_bool $a $b }.
-interpretation "or_bool" 'or_bool x y = 
- (cic:/matita/freescale/extra/or_bool.con x y).
+interpretation "or_bool" 'or_bool x y = (or_bool x y).
 
 (* \odot *)
 notation "hvbox(a break ⊙ b)" left associative with precedence 33
  for @{ 'xor_bool $a $b }.
-interpretation "xor_bool" 'xor_bool x y = 
- (cic:/matita/freescale/extra/xor_bool.con x y).
+interpretation "xor_bool" 'xor_bool x y = (xor_bool x y).
 
 (* ProdT e' gia' definito, aggiungo Prod3T e Prod4T e Prod5T *)
 
@@ -147,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".
@@ -192,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.