]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/Z/compare.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / Z / compare.ma
index 4ae6dd2c41e636921d45a8b040432baeb7c0237c..4a5025975fe3d8d4b3f610563b565dce39be6bac 100644 (file)
@@ -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,36 +41,39 @@ theorem eqZb_to_Prop:
 \forall x,y:Z. 
 match eqZb x y with
 [ true \Rightarrow x=y
-| false \Rightarrow \lnot x=y].
-intros.elim x.
-elim y.
-simplify.reflexivity.
-simplify.apply not_eq_OZ_neg.
-simplify.apply not_eq_OZ_pos.
-elim y.
-simplify.intro.apply not_eq_OZ_neg n ?.apply sym_eq.assumption.
-simplify.apply eqb_elim.intro.simplify.apply eq_f.assumption.
-intro.simplify.intro.apply H.apply inj_neg.assumption.
-simplify.intro.apply not_eq_pos_neg n1 n ?.apply sym_eq.assumption.
-elim y.
-simplify.intro.apply not_eq_OZ_pos n ?.apply sym_eq.assumption.
-simplify.apply not_eq_pos_neg.
-simplify.apply eqb_elim.intro.simplify.apply eq_f.assumption.
-intro.simplify.intro.apply H.apply inj_pos.assumption.
+| false \Rightarrow x \neq y].
+intros.
+elim x.
+  elim y.
+    simplify.reflexivity.
+    simplify.apply not_eq_OZ_pos.
+    simplify.apply not_eq_OZ_neg.
+  elim y.
+    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.unfold Not.intro.apply H.apply inj_pos.assumption.
+    simplify.apply not_eq_pos_neg.
+  elim y.
+    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.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
@@ -100,40 +101,43 @@ theorem Z_compare_to_Prop :
 | EQ \Rightarrow x=y
 | GT \Rightarrow y < x]. 
 intros.
-elim x. elim y.
-simplify.apply refl_eq.
-simplify.exact I.
-simplify.exact I.
-elim y. simplify.exact I.
-simplify. 
-cut match (nat_compare n1 n) with
-[ LT \Rightarrow n1<n
-| EQ \Rightarrow n1=n
-| GT \Rightarrow n<n1] \to 
-match (nat_compare n1 n) with
-[ LT \Rightarrow (S n1) \leq n
-| EQ \Rightarrow neg n = neg n1
-| GT \Rightarrow (S n) \leq n1]. 
-apply Hcut. apply nat_compare_to_Prop. 
-elim (nat_compare n1 n).
-simplify.exact H.
-simplify.exact H.
-simplify.apply eq_f.apply sym_eq.exact H.
-simplify.exact I.
-elim y.simplify.exact I.
-simplify.exact I.
-simplify.
-cut match (nat_compare n n1) with
-[ LT \Rightarrow n<n1
-| EQ \Rightarrow n=n1
-| GT \Rightarrow n1<n] \to 
-match (nat_compare n n1) with
-[ LT \Rightarrow (S n) \leq n1
-| EQ \Rightarrow pos n = pos n1
-| GT \Rightarrow (S n1) \leq n]. 
-apply Hcut. apply nat_compare_to_Prop. 
-elim (nat_compare n n1).
-simplify.exact H.
-simplify.exact H.
-simplify.apply eq_f.exact H.
+elim x. 
+  elim y.
+    simplify.apply refl_eq.
+    simplify.exact I.
+    simplify.exact I.
+  elim y.
+    simplify.exact I.
+    simplify.
+      cut (match (nat_compare n n1) with
+      [ LT \Rightarrow n<n1
+      | EQ \Rightarrow n=n1
+      | GT \Rightarrow n1<n] \to 
+      match (nat_compare n n1) with
+      [ LT \Rightarrow (S n) \leq n1
+      | EQ \Rightarrow pos n = pos n1
+      | GT \Rightarrow (S n1) \leq n]). 
+        apply Hcut.apply nat_compare_to_Prop. 
+        elim (nat_compare n n1).
+          simplify.exact H.
+          simplify.apply eq_f.exact H.
+          simplify.exact H.
+    simplify.exact I.    
+  elim y. 
+    simplify.exact I.
+    simplify.exact I.
+    simplify. 
+      cut (match (nat_compare n1 n) with
+      [ LT \Rightarrow n1<n
+      | EQ \Rightarrow n1=n
+      | GT \Rightarrow n<n1] \to 
+      match (nat_compare n1 n) with
+      [ LT \Rightarrow (S n1) \leq n
+      | EQ \Rightarrow neg n = neg n1
+      | GT \Rightarrow (S n) \leq n1]). 
+        apply Hcut. apply nat_compare_to_Prop. 
+        elim (nat_compare n1 n).
+          simplify.exact H.
+          simplify.apply eq_f.apply sym_eq.exact H.
+          simplify.exact H.
 qed.