X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Fle_arith.ma;h=a222d36bab2df5b26df3782439656478c936caae;hb=4dc47c9675ffd5fa50296ffaa9b5997501518c98;hp=ae386d313882f2b37903ae28c60c77dfa9fda072;hpb=f275a4daa42f8ab455020ec5a7ce7bf69b460c37;p=helm.git diff --git a/helm/software/matita/library/nat/le_arith.ma b/helm/software/matita/library/nat/le_arith.ma index ae386d313..a222d36ba 100644 --- a/helm/software/matita/library/nat/le_arith.ma +++ b/helm/software/matita/library/nat/le_arith.ma @@ -12,17 +12,16 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/nat/le_arith". - include "nat/times.ma". include "nat/orders.ma". (* plus *) theorem monotonic_le_plus_r: \forall n:nat.monotonic nat le (\lambda m.n + m). -simplify.intros.elim n. -simplify.assumption. -simplify.apply le_S_S.assumption. +simplify.intros.elim n + [simplify.assumption. + |simplify.apply le_S_S.assumption + ] qed. theorem le_plus_r: \forall p,n,m:nat. n \le m \to p + n \le p + m @@ -41,9 +40,21 @@ 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. +(** +auto. +*) +apply (transitive_le (plus n1 m1) (plus n1 m2) (plus n2 m2) ? ?); + [apply (monotonic_le_plus_r n1 m1 m2 ?). + apply (H1). + |apply (monotonic_le_plus_l m2 n1 n2 ?). + apply (H). + ] +(* end auto($Revision$) proof: TIME=0.61 SIZE=100 DEPTH=100 *) +(* 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. @@ -62,6 +73,16 @@ rewrite < sym_plus. apply le_plus_n. qed. +theorem le_plus_to_le: +\forall a,n,m. a + n \le a + m \to n \le m. +intro. +elim a + [assumption + |apply H. + apply le_S_S_to_le.assumption + ] +qed. + (* times *) theorem monotonic_le_times_r: \forall n:nat.monotonic nat le (\lambda m. n * m). @@ -98,3 +119,50 @@ intros.elim H.simplify. elim (plus_n_O ?).apply le_n. simplify.rewrite < sym_plus.apply le_plus_n. qed. + +theorem le_times_to_le: +\forall a,n,m. S O \le a \to a * n \le a * m \to n \le m. +intro. +apply nat_elim2;intros + [apply le_O_n + |apply False_ind. + rewrite < times_n_O in H1. + generalize in match H1. + apply (lt_O_n_elim ? H). + intros. + simplify in H2. + apply (le_to_not_lt ? ? H2). + apply lt_O_S + |apply le_S_S. + apply H + [assumption + |rewrite < times_n_Sm in H2. + rewrite < times_n_Sm in H2. + apply (le_plus_to_le a). + assumption + ] + ] +qed. + +theorem le_S_times_SSO: \forall n,m.O < m \to +n \le m \to S n \le (S(S O))*m. +intros. +simplify. +rewrite > plus_n_O. +simplify.rewrite > plus_n_Sm. +apply le_plus + [assumption + |rewrite < plus_n_O. + assumption + ] +qed. +(*0 and times *) +theorem O_lt_const_to_le_times_const: \forall a,c:nat. +O \lt c \to a \le a*c. +intros. +rewrite > (times_n_SO a) in \vdash (? % ?). +apply le_times +[ apply le_n +| assumption +] +qed.