]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/Z/z.ma
added a function to reorder the metasenv.
[helm.git] / helm / matita / library / Z / z.ma
index a5e0fba8f97271cc3102a686631d71e26e5be71b..207457a913e4e9c91574cf2e630ea61f3f2feb93 100644 (file)
@@ -54,18 +54,10 @@ match OZ_test z with
 |false \Rightarrow \lnot (z=OZ)].
 intros.elim z.
 simplify.reflexivity.
-simplify.intros.
-cut match neg n with 
-[ OZ \Rightarrow True 
-| (pos n) \Rightarrow False
-| (neg n) \Rightarrow False].
-apply Hcut.rewrite > H.simplify.exact I.
-simplify.intros.
-cut match pos n with 
-[ OZ \Rightarrow True 
-| (pos n) \Rightarrow False
-| (neg n) \Rightarrow False].
-apply Hcut. rewrite > H.simplify.exact I.
+simplify.intros [H].
+discriminate H.
+simplify.intros [H].
+discriminate H.
 qed.
 
 (* discrimination *)
@@ -90,56 +82,53 @@ 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. \lnot (OZ = (pos n)).
-simplify.intros.
-change with
-  match pos n with
-  [ OZ \Rightarrow True
-  | (pos n) \Rightarrow False
-  | (neg n) \Rightarrow False].
-rewrite < H.
-simplify.exact I.
+simplify.intros [n; H].
+discriminate H.
 qed.
 
 theorem not_eq_OZ_neg :\forall n:nat. \lnot (OZ = (neg n)).
-simplify.intros.
-change with
-  match neg n with
-  [ OZ \Rightarrow True
-  | (pos n) \Rightarrow False
-  | (neg n) \Rightarrow False].
-rewrite < H.
-simplify.exact I.
+simplify.intros [n; H].
+discriminate H.
 qed.
 
 theorem not_eq_pos_neg :\forall n,m:nat. \lnot ((pos n) = (neg m)).
-simplify.intros.
-change with
-  match neg m with
-  [ OZ \Rightarrow False
-  | (pos n) \Rightarrow True
-  | (neg n) \Rightarrow False].
-rewrite < H.
-simplify.exact I.
+simplify.intros [n; m; H].
+discriminate H.
 qed.
 
 theorem decidable_eq_Z : \forall x,y:Z. decidable (x=y).
 intros.simplify.
-elim x.elim y.
-left.reflexivity.
-right.apply not_eq_OZ_neg.
-right.apply not_eq_OZ_pos.
-elim y.right.intro.
-apply not_eq_OZ_neg n ?.apply sym_eq.assumption.
-elim (decidable_eq_nat n n1:(Or (n=n1) ((n=n1) \to False))).
-left.apply eq_f.assumption.
-right.intro.apply H.apply injective_neg.assumption.
-right.intro.apply not_eq_pos_neg n1 n ?.apply sym_eq.assumption.
-elim y.right.intro.
-apply not_eq_OZ_pos n ?.apply sym_eq.assumption.
-right.apply not_eq_pos_neg.
-elim (decidable_eq_nat n n1:(Or (n=n1) ((n=n1) \to False))).
-left.apply eq_f.assumption.
-right.intro.apply H.apply injective_pos.assumption.
+elim x.
+(* goal: x=OZ *)
+  elim y.
+  (* goal: x=OZ y=OZ *)
+    left.reflexivity.
+  (* goal: x=OZ 2=2 *)
+    right.apply not_eq_OZ_pos.
+  (* goal: x=OZ 2=3 *)
+    right.apply not_eq_OZ_neg.
+(* goal: x=pos *)
+  elim y.
+  (* goal: x=pos y=OZ *)
+    right.intro.
+    apply not_eq_OZ_pos n. symmetry. assumption.
+  (* goal: x=pos y=pos *)
+    elim (decidable_eq_nat n n1:(Or (n=n1) ((n=n1) \to False))).
+    left.apply eq_f.assumption.
+    right.intros [H_inj].apply H. injection H_inj. assumption.
+  (* goal: x=pos y=neg *)
+    right.intro.apply not_eq_pos_neg n n1. assumption.
+(* goal: x=neg *)
+  elim y.
+  (* goal: x=neg y=OZ *)
+    right.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.
+  (* goal: x=neg y=neg *)
+    elim (decidable_eq_nat n n1:(Or (n=n1) ((n=n1) \to False))).
+    left.apply eq_f.assumption.
+    right.intro.apply H.apply injective_neg.assumption.
 qed.
 
 (* end discrimination *)
@@ -163,16 +152,22 @@ definition Zpred \def
 | (neg n) \Rightarrow neg (S n)].
 
 theorem Zpred_Zsucc: \forall z:Z. Zpred (Zsucc z) = z.
-intros.elim z.reflexivity.
-elim n.reflexivity.
-reflexivity.
-reflexivity.
+intros.
+elim z.
+  reflexivity.
+  reflexivity.
+  elim n.
+    reflexivity.
+    reflexivity.
 qed.
 
 theorem Zsucc_Zpred: \forall z:Z. Zsucc (Zpred z) = z.
-intros.elim z.reflexivity.
-reflexivity.
-elim n.reflexivity.
-reflexivity.
+intros.
+elim z.
+  reflexivity.
+  elim n.
+    reflexivity.
+    reflexivity.
+  reflexivity.
 qed.