]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/le_arith.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / nat / le_arith.ma
index 7ece3fdc75b7053e1413bc30f963264d086cbf3b..a76183063b1e6d211d5c03a75cd1845c844c73aa 100644 (file)
@@ -31,7 +31,7 @@ theorem le_plus_r: \forall p,n,m:nat. n \le m \to p + n \le p + m
 theorem monotonic_le_plus_l: 
 \forall m:nat.monotonic nat le (\lambda n.n + m).
 simplify.intros.
-rewrite < sym_plus.rewrite < sym_plus m.
+rewrite < sym_plus.rewrite < (sym_plus m).
 apply le_plus_r.assumption.
 qed.
 
@@ -41,13 +41,13 @@ theorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
 theorem le_plus: \forall n1,n2,m1,m2:nat. n1 \le n2  \to m1 \le m2 
 \to n1 + m1 \le n2 + m2.
 intros.
-apply trans_le ? (n2 + m1).
+apply (trans_le ? (n2 + m1)).
 apply le_plus_l.assumption.
 apply le_plus_r.assumption.
 qed.
 
 theorem le_plus_n :\forall n,m:nat. m \le n + m.
-intros.change with O+m \le n+m.
+intros.change with (O+m \le n+m).
 apply le_plus_l.apply le_O_n.
 qed.
 
@@ -73,7 +73,7 @@ theorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
 theorem monotonic_le_times_l: 
 \forall m:nat.monotonic nat le (\lambda n.n*m).
 simplify.intros.
-rewrite < sym_times.rewrite < sym_times m.
+rewrite < sym_times.rewrite < (sym_times m).
 apply le_times_r.assumption.
 qed.
 
@@ -83,7 +83,7 @@ theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
 theorem le_times: \forall n1,n2,m1,m2:nat. n1 \le n2  \to m1 \le m2 
 \to n1*m1 \le n2*m2.
 intros.
-apply trans_le ? (n2*m1).
+apply (trans_le ? (n2*m1)).
 apply le_times_l.assumption.
 apply le_times_r.assumption.
 qed.
@@ -93,4 +93,3 @@ intros.elim H.simplify.
 elim (plus_n_O ?).apply le_n.
 simplify.rewrite < sym_plus.apply le_plus_n.
 qed.
-