]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/assembly/freescale/extra.ma
fixed some notational collisions
[helm.git] / helm / software / matita / contribs / assembly / freescale / extra.ma
index 94f4a47547eebf1a1048c94474c14f35e4e90625..25e0cb6fab5fe0f00cb7edf9d5d6419d9ce63fc3 100644 (file)
@@ -54,29 +54,35 @@ definition eq_bool ≝
  [ true ⇒ b2
  | false ⇒ not_bool b2 ].
 
+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.
+
 (* \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 *)