X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Flibrary%2Fnat%2Fle_arith.ma;h=90f3e182b2f51441b34c06054816a6c1cdcc9a82;hb=562e9adf40098e11d8f0bc2711a7f665360c2231;hp=a76183063b1e6d211d5c03a75cd1845c844c73aa;hpb=7f2444c2670cadafddd8785b687ef312158376b0;p=helm.git diff --git a/matita/library/nat/le_arith.ma b/matita/library/nat/le_arith.ma index a76183063..90f3e182b 100644 --- a/matita/library/nat/le_arith.ma +++ b/matita/library/nat/le_arith.ma @@ -51,12 +51,27 @@ intros.change with (O+m \le n+m). apply le_plus_l.apply le_O_n. qed. +theorem le_plus_n_r :\forall n,m:nat. m \le m + n. +intros.rewrite > sym_plus. +apply le_plus_n. +qed. + theorem eq_plus_to_le: \forall n,m,p:nat.n=m+p \to m \le n. intros.rewrite > H. 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). @@ -93,3 +108,27 @@ 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. \ No newline at end of file