]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/assembly/freescale/extra.ma
Semantic analysis implemented (sort of).
[helm.git] / helm / software / matita / contribs / assembly / freescale / extra.ma
index bc76c4d6b7f0addf954b61fa572c3ad3657d803e..25e0cb6fab5fe0f00cb7edf9d5d6419d9ce63fc3 100644 (file)
@@ -54,6 +54,16 @@ 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 }.