X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2FZ%2Fz.ma;h=d18c80b23ac323b996280a835a52912fc2ffc8ce;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=6ba305d98c05d2b8b50468f47dee83f5b8b41023;hpb=f263e4ec717d5ec2e7f9c057855f8223f81baae8;p=helm.git diff --git a/helm/matita/library/Z/z.ma b/helm/matita/library/Z/z.ma index 6ba305d98..d18c80b23 100644 --- a/helm/matita/library/Z/z.ma +++ b/helm/matita/library/Z/z.ma @@ -12,10 +12,10 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/Z/". +set "baseuri" "cic:/matita/Z/z". -include "nat/nat.ma". include "datatypes/bool.ma". +include "nat/nat.ma". inductive Z : Set \def OZ : Z @@ -34,37 +34,105 @@ definition neg_Z_of_nat \def [ O \Rightarrow OZ | (S n)\Rightarrow neg n]. -definition absZ \def +definition abs \def \lambda z. match z with [ OZ \Rightarrow O | (pos n) \Rightarrow n | (neg n) \Rightarrow n]. -definition OZ_testb \def +definition OZ_test \def \lambda z. match z with [ OZ \Rightarrow true | (pos n) \Rightarrow false | (neg n) \Rightarrow false]. -theorem OZ_discr : -\forall z. if_then_else (OZ_testb z) (eq Z z OZ) (Not (eq Z z OZ)). -intros.elim z.simplify.reflexivity. -simplify.intros. -cut match neg e1 with -[ OZ \Rightarrow True -| (pos n) \Rightarrow False -| (neg n) \Rightarrow False]. -apply Hcut.rewrite > H.simplify.exact I. -simplify.intros. -cut match pos e2 with -[ OZ \Rightarrow True -| (pos n) \Rightarrow False -| (neg n) \Rightarrow False]. -apply Hcut. rewrite > H.simplify.exact I. +theorem OZ_test_to_Prop :\forall z:Z. +match OZ_test z with +[true \Rightarrow z=OZ +|false \Rightarrow z \neq OZ]. +intros.elim z. +simplify.reflexivity. +simplify. unfold Not. intros (H). +discriminate H. +simplify. unfold Not. intros (H). +discriminate H. +qed. + +(* discrimination *) +theorem injective_pos: injective nat Z pos. +unfold injective. +intros. +change with (abs (pos x) = abs (pos y)). +apply eq_f.assumption. qed. +variant inj_pos : \forall n,m:nat. pos n = pos m \to n = m +\def injective_pos. + +theorem injective_neg: injective nat Z neg. +unfold injective. +intros. +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. OZ \neq pos n. +unfold Not.intros (n H). +discriminate H. +qed. + +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. 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.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 *) + definition Zsucc \def \lambda z. match z with [ OZ \Rightarrow pos O @@ -83,240 +151,23 @@ definition Zpred \def | (S p) \Rightarrow pos p] | (neg n) \Rightarrow neg (S n)]. -theorem Zpred_succ: \forall z:Z. eq Z (Zpred (Zsucc z)) z. -intros.elim z.reflexivity. -elim e1.reflexivity. -reflexivity. -reflexivity. -qed. - -theorem Zsucc_pred: \forall z:Z. eq Z (Zsucc (Zpred z)) z. -intros.elim z.reflexivity. -reflexivity. -elim e2.reflexivity. -reflexivity. -qed. - -let rec Zplus x y : Z \def - match x with - [ OZ \Rightarrow y - | (pos m) \Rightarrow - match y with - [ OZ \Rightarrow x - | (pos n) \Rightarrow (pos (S (plus m n))) - | (neg n) \Rightarrow - match nat_compare m n with - [ LT \Rightarrow (neg (pred (minus n m))) - | EQ \Rightarrow OZ - | GT \Rightarrow (pos (pred (minus m n)))]] - | (neg m) \Rightarrow - match y with - [ OZ \Rightarrow x - | (pos n) \Rightarrow - match nat_compare m n with - [ LT \Rightarrow (pos (pred (minus n m))) - | EQ \Rightarrow OZ - | GT \Rightarrow (neg (pred (minus m n)))] - | (neg n) \Rightarrow (neg (S (plus m n)))]]. - -theorem Zplus_z_O: \forall z:Z. eq Z (Zplus z OZ) z. -intro.elim z. -simplify.reflexivity. -simplify.reflexivity. -simplify.reflexivity. -qed. - -theorem sym_Zplus : \forall x,y:Z. eq Z (Zplus x y) (Zplus y x). -intros.elim x.simplify.rewrite > Zplus_z_O.reflexivity. -elim y.simplify.reflexivity. -simplify. -rewrite < sym_plus.reflexivity. -simplify. -rewrite > nat_compare_invert. -simplify.elim nat_compare ? ?.simplify.reflexivity. -simplify. reflexivity. -simplify. reflexivity. -elim y.simplify.reflexivity. -simplify.rewrite > nat_compare_invert. -simplify.elim nat_compare ? ?.simplify.reflexivity. -simplify. reflexivity. -simplify. reflexivity. -simplify.elim (sym_plus ? ?).reflexivity. -qed. - -theorem Zpred_neg : \forall z:Z. eq Z (Zpred z) (Zplus (neg O) z). -intros.elim z. -simplify.reflexivity. -simplify.reflexivity. -elim e2.simplify.reflexivity. -simplify.reflexivity. -qed. - -theorem Zsucc_pos : \forall z:Z. eq Z (Zsucc z) (Zplus (pos O) z). -intros.elim z. -simplify.reflexivity. -elim e1.simplify.reflexivity. -simplify.reflexivity. -simplify.reflexivity. -qed. - -theorem Zplus_succ_pred_pp : -\forall n,m. eq Z (Zplus (pos n) (pos m)) (Zplus (Zsucc (pos n)) (Zpred (pos m))). +theorem Zpred_Zsucc: \forall z:Z. Zpred (Zsucc z) = z. intros. -elim n.elim m. -simplify.reflexivity. -simplify.reflexivity. -elim m. -simplify. -rewrite < plus_n_O.reflexivity. -simplify. -rewrite < plus_n_Sm.reflexivity. +elim z. + reflexivity. + reflexivity. + elim n. + reflexivity. + reflexivity. qed. -theorem Zplus_succ_pred_pn : -\forall n,m. eq Z (Zplus (pos n) (neg m)) (Zplus (Zsucc (pos n)) (Zpred (neg m))). -intros.reflexivity. -qed. - -theorem Zplus_succ_pred_np : -\forall n,m. eq Z (Zplus (neg n) (pos m)) (Zplus (Zsucc (neg n)) (Zpred (pos m))). +theorem Zsucc_Zpred: \forall z:Z. Zsucc (Zpred z) = z. intros. -elim n.elim m. -simplify.reflexivity. -simplify.reflexivity. -elim m. -simplify.reflexivity. -simplify.reflexivity. -qed. - -theorem Zplus_succ_pred_nn: -\forall n,m. eq Z (Zplus (neg n) (neg m)) (Zplus (Zsucc (neg n)) (Zpred (neg m))). -intros. -elim n.elim m. -simplify.reflexivity. -simplify.reflexivity. -elim m. -simplify.rewrite < plus_n_Sm.reflexivity. -simplify.rewrite > plus_n_Sm.reflexivity. -qed. - -theorem Zplus_succ_pred: -\forall x,y. eq Z (Zplus x y) (Zplus (Zsucc x) (Zpred y)). -intros. -elim x. elim y. -simplify.reflexivity. -simplify.reflexivity. -rewrite < Zsucc_pos.rewrite > Zsucc_pred.reflexivity. -elim y.rewrite < sym_Zplus.rewrite < sym_Zplus (Zpred OZ). -rewrite < Zpred_neg.rewrite > Zpred_succ. -simplify.reflexivity. -rewrite < Zplus_succ_pred_nn.reflexivity. -apply Zplus_succ_pred_np. -elim y.simplify.reflexivity. -apply Zplus_succ_pred_pn. -apply Zplus_succ_pred_pp. -qed. - -theorem Zsucc_plus_pp : -\forall n,m. eq Z (Zplus (Zsucc (pos n)) (pos m)) (Zsucc (Zplus (pos n) (pos m))). -intros.reflexivity. -qed. - -theorem Zsucc_plus_pn : -\forall n,m. eq Z (Zplus (Zsucc (pos n)) (neg m)) (Zsucc (Zplus (pos n) (neg m))). -intros. -apply nat_double_ind -(\lambda n,m. eq Z (Zplus (Zsucc (pos n)) (neg m)) (Zsucc (Zplus (pos n) (neg m)))).intro. -intros.elim n1. -simplify. reflexivity. -elim e1.simplify. reflexivity. -simplify. reflexivity. -intros. elim n1. -simplify. reflexivity. -simplify.reflexivity. -intros. -rewrite < (Zplus_succ_pred_pn ? m1). -elim H.reflexivity. -qed. - -theorem Zsucc_plus_nn : -\forall n,m. eq Z (Zplus (Zsucc (neg n)) (neg m)) (Zsucc (Zplus (neg n) (neg m))). -intros. -apply nat_double_ind -(\lambda n,m. eq Z (Zplus (Zsucc (neg n)) (neg m)) (Zsucc (Zplus (neg n) (neg m)))).intro. -intros.elim n1. -simplify. reflexivity. -elim e1.simplify. reflexivity. -simplify. reflexivity. -intros. elim n1. -simplify. reflexivity. -simplify.reflexivity. -intros. -rewrite < (Zplus_succ_pred_nn ? m1). -reflexivity. +elim z. + reflexivity. + elim n. + reflexivity. + reflexivity. + reflexivity. qed. -theorem Zsucc_plus_np : -\forall n,m. eq Z (Zplus (Zsucc (neg n)) (pos m)) (Zsucc (Zplus (neg n) (pos m))). -intros. -apply nat_double_ind -(\lambda n,m. eq Z (Zplus (Zsucc (neg n)) (pos m)) (Zsucc (Zplus (neg n) (pos m)))). -intros.elim n1. -simplify. reflexivity. -elim e1.simplify. reflexivity. -simplify. reflexivity. -intros. elim n1. -simplify. reflexivity. -simplify.reflexivity. -intros. -rewrite < H. -rewrite < (Zplus_succ_pred_np ? (S m1)). -reflexivity. -qed. - - -theorem Zsucc_plus : \forall x,y:Z. eq Z (Zplus (Zsucc x) y) (Zsucc (Zplus x y)). -intros.elim x.elim y. -simplify. reflexivity. -rewrite < Zsucc_pos.reflexivity. -simplify.reflexivity. -elim y.rewrite < sym_Zplus.rewrite < sym_Zplus OZ.simplify.reflexivity. -apply Zsucc_plus_nn. -apply Zsucc_plus_np. -elim y. -rewrite < sym_Zplus OZ.reflexivity. -apply Zsucc_plus_pn. -apply Zsucc_plus_pp. -qed. - -theorem Zpred_plus : \forall x,y:Z. eq Z (Zplus (Zpred x) y) (Zpred (Zplus x y)). -intros. -cut eq Z (Zpred (Zplus x y)) (Zpred (Zplus (Zsucc (Zpred x)) y)). -rewrite > Hcut. -rewrite > Zsucc_plus. -rewrite > Zpred_succ. -reflexivity. -rewrite > Zsucc_pred. -reflexivity. -qed. - -theorem assoc_Zplus : -\forall x,y,z:Z. eq Z (Zplus x (Zplus y z)) (Zplus (Zplus x y) z). -intros.elim x.simplify.reflexivity. -elim e1.rewrite < (Zpred_neg (Zplus y z)). -rewrite < (Zpred_neg y). -rewrite < Zpred_plus. -reflexivity. -rewrite > Zpred_plus (neg e). -rewrite > Zpred_plus (neg e). -rewrite > Zpred_plus (Zplus (neg e) y). -apply f_equal.assumption. -elim e2.rewrite < Zsucc_pos. -rewrite < Zsucc_pos. -rewrite > Zsucc_plus. -reflexivity. -rewrite > Zsucc_plus (pos e1). -rewrite > Zsucc_plus (pos e1). -rewrite > Zsucc_plus (Zplus (pos e1) y). -apply f_equal.assumption. -qed.