]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/num/quatern_lemmas.ma
mod change (-x)
[helm.git] / helm / software / matita / contribs / ng_assembly / num / quatern_lemmas.ma
old mode 100755 (executable)
new mode 100644 (file)
index c2c8a33..a3ff945
@@ -16,7 +16,7 @@
 (*                          Progetto FreeScale                            *)
 (*                                                                        *)
 (*   Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it              *)
-(*   Ultima modifica: 05/08/2009                                          *)
+(*   Sviluppo: 2008-2010                                                  *)
 (*                                                                        *)
 (* ********************************************************************** *)
 
@@ -27,6 +27,7 @@ include "num/bool_lemmas.ma".
 (* QUATERNARI *)
 (* ********** *)
 
+(*
 ndefinition quatern_destruct_aux ≝
 Πn1,n2:quatern.ΠP:Prop.n1 = n2 →
  match eq_qu n1 n2 with [ true ⇒ P → P | false ⇒ P ].
@@ -38,6 +39,7 @@ ndefinition quatern_destruct : quatern_destruct_aux.
  nnormalize;
  napply (λx.x).
 nqed.
+*)
 
 nlemma eq_to_eqqu : ∀n1,n2.n1 = n2 → eq_qu n1 n2 = true.
  #n1; #n2; #H;
@@ -61,7 +63,7 @@ nlemma eqqu_to_eq : ∀n1,n2.eq_qu n1 n2 = true → n1 = n2.
  ncases n2;
  nnormalize;
  ##[ ##1,6,11,16: #H; napply refl_eq
- ##| ##*: #H; napply (bool_destruct … H)
+ ##| ##*: #H; ndestruct (*napply (bool_destruct … H)*)
  ##]
 nqed.