]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/compare.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / nat / compare.ma
index 6431f5f293e219fcc5af7ebe84c315c4da5d0ae4..2647315804661a23db188e36b4955c0b01c5152a 100644 (file)
@@ -42,13 +42,13 @@ intro.elim n1.
 simplify.reflexivity.
 simplify.apply not_eq_O_S.
 intro.
-simplify.
+simplify.unfold Not.
 intro. apply (not_eq_O_S n1).apply sym_eq.assumption.
 intros.simplify.
 generalize in match H.
 elim ((eqb n1 m1)).
 simplify.apply eq_f.apply H1.
-simplify.intro.apply H1.apply inj_S.assumption.
+simplify.unfold Not.intro.apply H1.apply inj_S.assumption.
 qed.
 
 theorem eqb_elim : \forall n,m:nat.\forall P:bool \to Prop.
@@ -69,6 +69,28 @@ intro.elim n.simplify.reflexivity.
 simplify.assumption.
 qed.
 
+theorem eqb_true_to_eq: \forall n,m:nat.
+eqb n m = true \to n = m.
+intros.
+change with 
+match true with
+[ true  \Rightarrow n = m 
+| false \Rightarrow n \neq m].
+rewrite < H.
+apply eqb_to_Prop. 
+qed.
+
+theorem eqb_false_to_not_eq: \forall n,m:nat.
+eqb n m = false \to n \neq m.
+intros.
+change with 
+match false with
+[ true  \Rightarrow n = m 
+| false \Rightarrow n \neq m].
+rewrite < H.
+apply eqb_to_Prop. 
+qed.
+
 theorem eq_to_eqb_true: \forall n,m:nat.
 n = m \to eqb n m = true.
 intros.apply (eqb_elim n m).
@@ -104,7 +126,7 @@ simplify.exact le_O_n.
 simplify.exact not_le_Sn_O.
 intros 2.simplify.elim ((leb n1 m1)).
 simplify.apply le_S_S.apply H.
-simplify.intros.apply H.apply le_S_S_to_le.assumption.
+simplify.unfold Not.intros.apply H.apply le_S_S_to_le.assumption.
 qed.
 
 theorem leb_elim: \forall n,m:nat. \forall P:bool \to Prop. 
@@ -169,12 +191,12 @@ apply (nat_elim2 (\lambda n,m.match (nat_compare n m) with
   | EQ \Rightarrow n=m
   | GT \Rightarrow m < n ])).
 intro.elim n1.simplify.reflexivity.
-simplify.apply le_S_S.apply le_O_n.
-intro.simplify.apply le_S_S. apply le_O_n.
+simplify.unfold lt.apply le_S_S.apply le_O_n.
+intro.simplify.unfold lt.apply le_S_S. apply le_O_n.
 intros 2.simplify.elim ((nat_compare n1 m1)).
-simplify. apply le_S_S.apply H.
+simplify. unfold lt. apply le_S_S.apply H.
 simplify. apply eq_f. apply H.
-simplify. apply le_S_S.apply H.
+simplify. unfold lt.apply le_S_S.apply H.
 qed.
 
 theorem nat_compare_n_m_m_n: \forall n,m:nat.