X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2FZ%2Fcompare.ma;h=4a5025975fe3d8d4b3f610563b565dce39be6bac;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=a8f8aca3918f12c1a7a60fb475338a81da32dfbe;hpb=78044035b4419e569df0d7f6a7f96fa32d21a19d;p=helm.git diff --git a/helm/matita/library/Z/compare.ma b/helm/matita/library/Z/compare.ma index a8f8aca39..4a5025975 100644 --- a/helm/matita/library/Z/compare.ma +++ b/helm/matita/library/Z/compare.ma @@ -14,8 +14,6 @@ set "baseuri" "cic:/matita/Z/compare". -include "datatypes/bool.ma". -include "datatypes/compare.ma". include "Z/orders.ma". include "nat/compare.ma". @@ -43,7 +41,7 @@ theorem eqZb_to_Prop: \forall x,y:Z. match eqZb x y with [ true \Rightarrow x=y -| false \Rightarrow \lnot x=y]. +| false \Rightarrow x \neq y]. intros. elim x. elim y. @@ -51,31 +49,31 @@ elim x. simplify.apply not_eq_OZ_pos. simplify.apply not_eq_OZ_neg. elim y. - simplify.intro.apply not_eq_OZ_pos n.apply sym_eq.assumption. + simplify.unfold Not.intro.apply (not_eq_OZ_pos n).apply sym_eq.assumption. simplify.apply eqb_elim. intro.simplify.apply eq_f.assumption. - intro.simplify.intro.apply H.apply inj_pos.assumption. + intro.simplify.unfold Not.intro.apply H.apply inj_pos.assumption. simplify.apply not_eq_pos_neg. elim y. - simplify.intro.apply not_eq_OZ_neg n.apply sym_eq.assumption. - simplify.intro.apply not_eq_pos_neg n1 n.apply sym_eq.assumption. + simplify.unfold Not.intro.apply (not_eq_OZ_neg n).apply sym_eq.assumption. + simplify.unfold Not.intro.apply (not_eq_pos_neg n1 n).apply sym_eq.assumption. simplify.apply eqb_elim. intro.simplify.apply eq_f.assumption. - intro.simplify.intro.apply H.apply inj_neg.assumption. + intro.simplify.unfold Not.intro.apply H.apply inj_neg.assumption. qed. theorem eqZb_elim: \forall x,y:Z.\forall P:bool \to Prop. -(x=y \to (P true)) \to (\lnot x=y \to (P false)) \to P (eqZb x y). +(x=y \to (P true)) \to (x \neq y \to (P false)) \to P (eqZb x y). intros. cut -match (eqZb x y) with +(match (eqZb x y) with [ true \Rightarrow x=y -| false \Rightarrow \lnot x=y] \to P (eqZb x y). +| false \Rightarrow x \neq y] \to P (eqZb x y)). apply Hcut. apply eqZb_to_Prop. elim (eqZb). -apply H H2. -apply H1 H2. +apply (H H2). +apply (H1 H2). qed. definition Z_compare : Z \to Z \to compare \def @@ -111,14 +109,14 @@ elim x. elim y. simplify.exact I. simplify. - cut match (nat_compare n n1) with + cut (match (nat_compare n n1) with [ LT \Rightarrow n