X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Flog.ma;h=39d9b2ecdfae7ecc55209fd8e979581fd28dbbc9;hb=10f29fdd78ee089a9a94446207b543d33d6c851c;hp=53200b1613204457774133d2d9fdaad71527918e;hpb=1f136ddf1cf133efbf10b72b8800e0f8a93b6f68;p=helm.git diff --git a/helm/software/matita/library/nat/log.ma b/helm/software/matita/library/nat/log.ma index 53200b161..39d9b2ecd 100644 --- a/helm/software/matita/library/nat/log.ma +++ b/helm/software/matita/library/nat/log.ma @@ -159,6 +159,50 @@ cases n ] qed. +theorem log_times_l: \forall p,n,m.O < n \to O < m \to S O < p \to +log p n+log p m \le log p (n*m) . +intros. +unfold log in ⊢ (? ? (% ? ?)). +apply f_m_to_le_max + [elim H + [rewrite > log_SO + [simplify. + rewrite < plus_n_O. + apply le_log_n_n. + assumption + |assumption + ] + |elim H1 + [rewrite > log_SO + [rewrite < plus_n_O. + rewrite < times_n_SO. + apply le_log_n_n. + assumption + |assumption + ] + |apply (trans_le ? (S n1 + S n2)) + [apply le_plus;apply le_log_n_n;assumption + |simplify. + apply le_S_S. + rewrite < plus_n_Sm. + change in ⊢ (? % ?) with ((S n1)+n2). + rewrite > sym_plus. + apply le_plus_r. + change with (n1 < n1*S n2). + rewrite > times_n_SO in ⊢ (? % ?). + apply lt_times_r1 + [assumption + |apply le_S_S.assumption + ] + ] + ] + ] + |apply le_to_leb_true. + rewrite > exp_plus_times. + apply le_times;apply le_exp_log;assumption + ] +qed. + theorem log_exp: \forall p,n,m.S O < p \to O < m \to log p ((exp p n)*m)=n+log p m. intros. @@ -286,6 +330,30 @@ intros.elim H1 [constructor 1 |apply (trans_le ? ? ? H3);apply le_log_plus;assumption] qed. + +theorem log_div: \forall p,n,m. S O < p \to O < m \to m \le n \to +log p (n/m) \le log p n -log p m. +intros. +apply le_plus_to_minus_r. +apply (trans_le ? (log p ((n/m)*m))) + [apply log_times_l + [apply le_times_to_le_div + [assumption + |rewrite < times_n_SO. + assumption + ] + |assumption + |assumption + ] + |apply le_log + [assumption + |rewrite > (div_mod n m) in ⊢ (? ? %) + [apply le_plus_n_r + |assumption + ] + ] + ] +qed. theorem log_n_n: \forall n. S O < n \to log n n = S O. intros.