]> matita.cs.unibo.it Git - helm.git/commitdiff
{discriminate,injection} => destruct
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Tue, 26 Sep 2006 08:26:17 +0000 (08:26 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Tue, 26 Sep 2006 08:26:17 +0000 (08:26 +0000)
matita/library/Z/z.ma
matita/library/list/list.ma

index 9532ea95deb4d3ffb2a7c66b1665f58aa7a6c462..d8b829ec1da6e60da84bf89fe20c94e80dc72351 100644 (file)
@@ -55,9 +55,9 @@ match OZ_test z with
 intros.elim z.
 simplify.reflexivity.
 simplify. unfold Not. intros (H).
-discriminate H.
+destruct H.
 simplify. unfold Not. intros (H).
-discriminate H.
+destruct H.
 qed.
 
 (* discrimination *)
@@ -85,17 +85,17 @@ variant inj_neg : \forall n,m:nat. neg n = neg m \to n = m
 
 theorem not_eq_OZ_pos: \forall n:nat. OZ \neq pos n.
 unfold Not.intros (n H).
-discriminate H.
+destruct H.
 qed.
 
 theorem not_eq_OZ_neg :\forall n:nat. OZ \neq neg n.
 unfold Not.intros (n H).
-discriminate H.
+destruct H.
 qed.
 
 theorem not_eq_pos_neg :\forall n,m:nat. pos n \neq neg m.
 unfold Not.intros (n m H).
-discriminate H.
+destruct H.
 qed.
 
 theorem decidable_eq_Z : \forall x,y:Z. decidable (x=y).
@@ -117,7 +117,7 @@ elim x.
   (* goal: x=pos y=pos *)
     elim (decidable_eq_nat n n1:((n=n1) \lor ((n=n1) \to False))).
     left.apply eq_f.assumption.
-    right.unfold Not.intros (H_inj).apply H. injection H_inj. assumption.
+    right.unfold Not.intros (H_inj).apply H. destruct H_inj. assumption.
   (* goal: x=pos y=neg *)
     right.unfold Not.intro.apply (not_eq_pos_neg n n1). assumption.
 (* goal: x=neg *)
index bb341d245fad82a43657e4e427956968fac8b3c2..9ecfd50e3ff41bf5cf89e6604ea2ca4986cf7e48 100644 (file)
@@ -44,7 +44,7 @@ theorem nil_cons:
   intros;
   unfold Not;
   intros;
-  discriminate H.
+  destruct H.
 qed.
 
 let rec id_list A (l: list A) on l :=