]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/Z/z.ma
new implementation of the destruct tactic,
[helm.git] / matita / library / Z / z.ma
index ea50a2cd9c3b70729df6c8c6bb4b16d661d80d41..ac530c38fe02b7df1766fd69ff93a637162131b0 100644 (file)
@@ -34,12 +34,17 @@ definition neg_Z_of_nat \def
 [ O \Rightarrow  OZ 
 | (S n)\Rightarrow  neg n].
 
+lemma pos_n_eq_S_n : \forall n : nat.
+  (pos n) = (S n).
+intro.reflexivity. 
+qed.
+
 definition abs \def
 \lambda z.
  match z with 
 [ OZ \Rightarrow O
-| (pos n) \Rightarrow n
-| (neg n) \Rightarrow n].
+| (pos n) \Rightarrow (S n)
+| (neg n) \Rightarrow (S n)].
 
 definition OZ_test \def
 \lambda z.
@@ -55,15 +60,16 @@ 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 *)
 theorem injective_pos: injective nat Z pos.
 unfold injective.
 intros.
+apply inj_S.
 change with (abs (pos x) = abs (pos y)).
 apply eq_f.assumption.
 qed.
@@ -74,6 +80,7 @@ variant inj_pos : \forall n,m:nat. pos n = pos m \to n = m
 theorem injective_neg: injective nat Z neg.
 unfold injective.
 intros.
+apply inj_S.
 change with (abs (neg x) = abs (neg y)).
 apply eq_f.assumption.
 qed.
@@ -83,17 +90,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).
@@ -115,7 +122,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. reflexivity.
   (* goal: x=pos y=neg *)
     right.unfold Not.intro.apply (not_eq_pos_neg n n1). assumption.
 (* goal: x=neg *)