X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2FZ%2Fz.ma;h=31c3a7003bdffca309780ce3e38122712efe4368;hb=6d297b12c480352eb2f156ab4515f73921ea2e81;hp=ea50a2cd9c3b70729df6c8c6bb4b16d661d80d41;hpb=55b82bd235d82ff7f0a40d980effe1efde1f5073;p=helm.git diff --git a/helm/software/matita/library/Z/z.ma b/helm/software/matita/library/Z/z.ma index ea50a2cd9..31c3a7003 100644 --- a/helm/software/matita/library/Z/z.ma +++ b/helm/software/matita/library/Z/z.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/Z/z". - include "datatypes/bool.ma". include "nat/nat.ma". @@ -22,24 +20,31 @@ inductive Z : Set \def | pos : nat \to Z | neg : nat \to Z. +interpretation "Integers" 'Z = Z. + definition Z_of_nat \def \lambda n. match n with [ O \Rightarrow OZ | (S n)\Rightarrow pos n]. -coercion cic:/matita/Z/z/Z_of_nat.con. +coercion Z_of_nat. definition neg_Z_of_nat \def \lambda n. match n with [ 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 *)