]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/Z/z.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / Z / z.ma
index d1a846da4754c07161cbaf668a895c0b22846096..d18c80b23ac323b996280a835a52912fc2ffc8ce 100644 (file)
@@ -54,15 +54,15 @@ match OZ_test z with
 |false \Rightarrow z \neq OZ].
 intros.elim z.
 simplify.reflexivity.
-simplify.intros (H).
+simplify. unfold Not. intros (H).
 discriminate H.
-simplify.intros (H).
+simplify. unfold Not. intros (H).
 discriminate H.
 qed.
 
 (* discrimination *)
 theorem injective_pos: injective nat Z pos.
-simplify.
+unfold injective.
 intros.
 change with (abs (pos x) = abs (pos y)).
 apply eq_f.assumption.
@@ -72,7 +72,7 @@ variant inj_pos : \forall n,m:nat. pos n = pos m \to n = m
 \def injective_pos.
 
 theorem injective_neg: injective nat Z neg.
-simplify.
+unfold injective.
 intros.
 change with (abs (neg x) = abs (neg y)).
 apply eq_f.assumption.
@@ -82,22 +82,22 @@ variant inj_neg : \forall n,m:nat. neg n = neg m \to n = m
 \def injective_neg.
 
 theorem not_eq_OZ_pos: \forall n:nat. OZ \neq pos n.
-simplify.intros (n H).
+unfold Not.intros (n H).
 discriminate H.
 qed.
 
 theorem not_eq_OZ_neg :\forall n:nat. OZ \neq neg n.
-simplify.intros (n H).
+unfold Not.intros (n H).
 discriminate H.
 qed.
 
 theorem not_eq_pos_neg :\forall n,m:nat. pos n \neq neg m.
-simplify.intros (n m H).
+unfold Not.intros (n m H).
 discriminate H.
 qed.
 
 theorem decidable_eq_Z : \forall x,y:Z. decidable (x=y).
-intros.simplify.
+intros.unfold decidable.
 elim x.
 (* goal: x=OZ *)
   elim y.
@@ -110,25 +110,25 @@ elim x.
 (* goal: x=pos *)
   elim y.
   (* goal: x=pos y=OZ *)
-    right.intro.
+    right.unfold Not.intro.
     apply (not_eq_OZ_pos n). symmetry. assumption.
   (* goal: x=pos y=pos *)
     elim (decidable_eq_nat n n1:((n=n1) \lor ((n=n1) \to False))).
     left.apply eq_f.assumption.
-    right.intros (H_inj).apply H. injection H_inj. assumption.
+    right.unfold Not.intros (H_inj).apply H. injection H_inj. assumption.
   (* goal: x=pos y=neg *)
-    right.intro.apply (not_eq_pos_neg n n1). assumption.
+    right.unfold Not.intro.apply (not_eq_pos_neg n n1). assumption.
 (* goal: x=neg *)
   elim y.
   (* goal: x=neg y=OZ *)
-    right.intro.
+    right.unfold Not.intro.
     apply (not_eq_OZ_neg n). symmetry. assumption.
   (* goal: x=neg y=pos *)
-    right. intro. apply (not_eq_pos_neg n1 n). symmetry. assumption.
+    right. unfold Not.intro. apply (not_eq_pos_neg n1 n). symmetry. assumption.
   (* goal: x=neg y=neg *)
     elim (decidable_eq_nat n n1:((n=n1) \lor ((n=n1) \to False))).
     left.apply eq_f.assumption.
-    right.intro.apply H.apply injective_neg.assumption.
+    right.unfold Not.intro.apply H.apply injective_neg.assumption.
 qed.
 
 (* end discrimination *)