X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2FZ%2Fz.ma;h=d18c80b23ac323b996280a835a52912fc2ffc8ce;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=a5e0fba8f97271cc3102a686631d71e26e5be71b;hpb=03fcee16d9c262aad38a47d0a409b684a965cc3f;p=helm.git diff --git a/helm/matita/library/Z/z.ma b/helm/matita/library/Z/z.ma index a5e0fba8f..d18c80b23 100644 --- a/helm/matita/library/Z/z.ma +++ b/helm/matita/library/Z/z.ma @@ -14,8 +14,8 @@ set "baseuri" "cic:/matita/Z/z". +include "datatypes/bool.ma". include "nat/nat.ma". -include "higher_order_defs/functions.ma". inductive Z : Set \def OZ : Z @@ -51,28 +51,20 @@ match z with theorem OZ_test_to_Prop :\forall z:Z. match OZ_test z with [true \Rightarrow z=OZ -|false \Rightarrow \lnot (z=OZ)]. +|false \Rightarrow z \neq 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. unfold Not. intros (H). +discriminate 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). +change with (abs (pos x) = abs (pos y)). apply eq_f.assumption. qed. @@ -80,66 +72,63 @@ 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). +change with (abs (neg x) = abs (neg y)). apply eq_f.assumption. qed. 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. +theorem not_eq_OZ_pos: \forall n:nat. OZ \neq pos n. +unfold Not.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. +theorem not_eq_OZ_neg :\forall n:nat. OZ \neq neg n. +unfold Not.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. +theorem not_eq_pos_neg :\forall n,m:nat. pos n \neq neg m. +unfold Not.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. +intros.unfold decidable. +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.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.unfold Not.intros (H_inj).apply H. injection H_inj. assumption. + (* goal: x=pos y=neg *) + right.unfold Not.intro.apply (not_eq_pos_neg n n1). assumption. +(* goal: x=neg *) + elim y. + (* goal: x=neg y=OZ *) + right.unfold Not.intro. + apply (not_eq_OZ_neg n). symmetry. assumption. + (* goal: x=neg y=pos *) + 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.unfold Not.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.