X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Flibrary%2Fnat%2Fle_arith.ma;h=9ab53fd7405a1946cf555974992c922dbdcc2acb;hb=e892d4cb9d19305ff88aa1b7dba6e3eaee41fd92;hp=ae386d313882f2b37903ae28c60c77dfa9fda072;hpb=03d30a78581b24842e1f425b41861e5ae8f912b5;p=helm.git diff --git a/matita/library/nat/le_arith.ma b/matita/library/nat/le_arith.ma index ae386d313..9ab53fd74 100644 --- a/matita/library/nat/le_arith.ma +++ b/matita/library/nat/le_arith.ma @@ -62,6 +62,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 +108,38 @@ 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. + +(*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.