]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/orders.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / nat / orders.ma
index 0ddd4b5c59571e69b8fe995151356085ef283604..6ec0c9992a68b18e15976e96e74ef853c2fc04eb 100644 (file)
@@ -54,7 +54,7 @@ interpretation "natural 'not greater than'" 'ngtr x y =
   (cic:/matita/logic/connectives/Not.con (cic:/matita/nat/orders/gt.con x y)).
 
 theorem transitive_le : transitive nat le.
-simplify.intros.elim H1.
+unfold transitive.intros.elim H1.
 assumption.
 apply le_S.assumption.
 qed.
@@ -63,7 +63,7 @@ theorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
 \def transitive_le.
 
 theorem transitive_lt: transitive nat lt.
-simplify.intros.elim H1.
+unfold transitive.unfold lt.intros.elim H1.
 apply le_S. assumption.
 apply le_S.assumption.
 qed.
@@ -105,11 +105,11 @@ qed.
 
 (* not le *)
 theorem not_le_Sn_O: \forall n:nat. S n \nleq O.
-intros.simplify.intros.apply (leS_to_not_zero ? ? H).
+intros.unfold Not.simplify.intros.apply (leS_to_not_zero ? ? H).
 qed.
 
 theorem not_le_Sn_n: \forall n:nat. S n \nleq n.
-intros.elim n.apply not_le_Sn_O.simplify.intros.cut (S n1 \leq n1).
+intros.elim n.apply not_le_Sn_O.unfold Not.simplify.intros.cut (S n1 \leq n1).
 apply H.assumption.
 apply le_S_S_to_le.assumption.
 qed.
@@ -119,19 +119,19 @@ theorem le_to_or_lt_eq : \forall n,m:nat.
 n \leq m \to n < m \lor n = m.
 intros.elim H.
 right.reflexivity.
-left.simplify.apply le_S_S.assumption.
+left.unfold lt.apply le_S_S.assumption.
 qed.
 
 (* not eq *)
 theorem lt_to_not_eq : \forall n,m:nat. n<m \to n \neq m.
-simplify.intros.cut ((le (S n) m) \to False).
+unfold Not.intros.cut ((le (S n) m) \to False).
 apply Hcut.assumption.rewrite < H1.
 apply not_le_Sn_n.
 qed.
 
 (* le vs. lt *)
 theorem lt_to_le : \forall n,m:nat. n<m \to n \leq m.
-simplify.intros.elim H.
+simplify.intros.unfold lt in H.elim H.
 apply le_S. apply le_n.
 apply le_S. assumption.
 qed.
@@ -145,13 +145,13 @@ theorem not_le_to_lt: \forall n,m:nat. n \nleq m \to m<n.
 intros 2.
 apply (nat_elim2 (\lambda n,m.n \nleq m \to m<n)).
 intros.apply (absurd (O \leq n1)).apply le_O_n.assumption.
-simplify.intros.apply le_S_S.apply le_O_n.
-simplify.intros.apply le_S_S.apply H.intros.apply H1.apply le_S_S.
+unfold Not.unfold lt.intros.apply le_S_S.apply le_O_n.
+unfold Not.unfold lt.intros.apply le_S_S.apply H.intros.apply H1.apply le_S_S.
 assumption.
 qed.
 
 theorem lt_to_not_le: \forall n,m:nat. n<m \to m \nleq n.
-simplify.intros 3.elim H.
+unfold Not.unfold lt.intros 3.elim H.
 apply (not_le_Sn_n n H1).
 apply H2.apply lt_to_le. apply H3.
 qed.
@@ -165,7 +165,7 @@ qed.
 theorem le_to_not_lt: \forall n,m:nat. le n m \to Not (lt m n).
 intros.
 change with (Not (le (S m) n)).
-apply lt_to_not_le.simplify.
+apply lt_to_not_le.unfold lt.
 apply le_S_S.assumption.
 qed.
 
@@ -195,12 +195,12 @@ qed.
 (* lt and le trans *)
 theorem lt_to_le_to_lt: \forall n,m,p:nat. lt n m \to le m p \to lt n p.
 intros.elim H1.
-assumption.simplify.apply le_S.assumption.
+assumption.unfold lt.apply le_S.assumption.
 qed.
 
 theorem le_to_lt_to_lt: \forall n,m,p:nat. le n m \to lt m p \to lt n p.
 intros 4.elim H.
-assumption.apply H2.simplify.
+assumption.apply H2.unfold lt.
 apply lt_to_le.assumption.
 qed.
 
@@ -217,7 +217,7 @@ qed.
 
 (* other abstract properties *)
 theorem antisymmetric_le : antisymmetric nat le.
-simplify.intros 2.
+unfold antisymmetric.intros 2.
 apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))).
 intros.apply le_n_O_to_eq.assumption.
 intros.apply False_ind.apply (not_le_Sn_O ? H).
@@ -232,11 +232,11 @@ theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
 theorem decidable_le: \forall n,m:nat. decidable (n \leq m).
 intros.
 apply (nat_elim2 (\lambda n,m.decidable (n \leq m))).
-intros.simplify.left.apply le_O_n.
-intros.simplify.right.exact (not_le_Sn_O n1).
-intros 2.simplify.intro.elim H.
+intros.unfold decidable.left.apply le_O_n.
+intros.unfold decidable.right.exact (not_le_Sn_O n1).
+intros 2.unfold decidable.intro.elim H.
 left.apply le_S_S.assumption.
-right.intro.apply H1.apply le_S_S_to_le.assumption.
+right.unfold Not.intro.apply H1.apply le_S_S_to_le.assumption.
 qed.
 
 theorem decidable_lt: \forall n,m:nat. decidable (n < m).
@@ -265,7 +265,7 @@ definition increasing \def \lambda f:nat \to nat.
 
 theorem increasing_to_monotonic: \forall f:nat \to nat.
 increasing f \to monotonic nat lt f.
-simplify.intros.elim H1.apply H.
+unfold monotonic.unfold lt.unfold increasing.unfold lt.intros.elim H1.apply H.
 apply (trans_le ? (f n1)).
 assumption.apply (trans_le ? (S (f n1))).
 apply le_n_Sn.
@@ -278,7 +278,7 @@ intros.elim n.
 apply le_O_n.
 apply (trans_le ? (S (f n1))).
 apply le_S_S.apply H1.
-simplify in H. apply H.
+simplify in H. unfold increasing in H.unfold lt in H.apply H.
 qed.
 
 theorem increasing_to_le: \forall f:nat \to nat. (increasing f) 
@@ -289,7 +289,7 @@ elim H1.
 apply (ex_intro ? ? (S a)).
 apply (trans_le ? (S (f a))).
 apply le_S_S.assumption.
-simplify in H.
+simplify in H.unfold increasing in H.unfold lt in H.
 apply H.
 qed.