X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Flog.ma;h=4c326ef25dcc19fe187ab98ef8e1cf04e1d584c6;hb=063523ae5f8da7e6458232f4afb6744ec86dc8bd;hp=d4e1b2bd991ce58ef77329a8e2bde2dfea7aad7a;hpb=947ee89dec9e60561dfac3ce7e1842f35f178cb8;p=helm.git diff --git a/helm/software/matita/library/nat/log.ma b/helm/software/matita/library/nat/log.ma index d4e1b2bd9..4c326ef25 100644 --- a/helm/software/matita/library/nat/log.ma +++ b/helm/software/matita/library/nat/log.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/nat/log". - include "datatypes/constructors.ma". include "nat/minimization.ma". include "nat/relevant_equations.ma". @@ -159,6 +157,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. @@ -271,20 +313,42 @@ apply (lt_exp_to_lt p) ] qed. -theorem le_log: \forall p,n,m. S O < p \to O < n \to n \le m \to +lemma le_log_plus: \forall p,n.S O < p \to log p n \leq log p (S n). +intros;apply (bool_elim ? (leb (p*(exp p n)) (S n))) + [simplify;intro;rewrite > H1;simplify;apply (trans_le ? n) + [apply le_log_n_n;assumption + |apply le_n_Sn] + |intro;unfold log;simplify;rewrite > H1;simplify;apply le_max_f_max_g; + intros;apply le_to_leb_true;constructor 2;apply leb_true_to_le;assumption] +qed. + +theorem le_log: \forall p,n,m. S O < p \to n \le m \to log p n \le log p m. +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_S_S_to_le. -apply (lt_exp_to_lt p) - [assumption - |apply (le_to_lt_to_lt ? n) - [apply le_exp_log. - assumption - |apply (le_to_lt_to_lt ? m) +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 - |apply lt_exp_log. + |rewrite < times_n_SO. assumption ] + |assumption + |assumption + ] + |apply le_log + [assumption + |rewrite > (div_mod n m) in ⊢ (? ? %) + [apply le_plus_n_r + |assumption + ] ] ] qed. @@ -331,14 +395,12 @@ apply antisymmetric_le [rewrite < (log_n_n i) in ⊢ (? % ?) [apply le_log [apply (trans_lt ? n);assumption - |apply (ltn_to_ltO ? ? H1) |apply le_n ] |apply (trans_lt ? n);assumption ] |apply le_log [apply (trans_lt ? n);assumption - |apply (ltn_to_ltO ? ? H1) |assumption ] ] @@ -350,6 +412,7 @@ intros.apply (lt_O_n_elim ? H).intros. simplify.reflexivity. qed. +(* theorem tech1: \forall n,i.O < n \to (exp (S n) (S(S i)))/(exp n (S i)) \le ((exp n i) + (exp (S n) (S i)))/(exp n i). intros. @@ -358,9 +421,8 @@ rewrite < eq_div_div_div_times [apply monotonic_div [apply lt_O_exp.assumption |apply le_S_S_to_le. - apply lt_times_to_lt_div - [assumption - |change in ⊢ (? % ?) with ((exp (S n) (S i)) + n*(exp (S n) (S i))). + apply lt_times_to_lt_div. + change in ⊢ (? % ?) with ((exp (S n) (S i)) + n*(exp (S n) (S i))). |apply (trans_le ? ((n)\sup(i)*(S n)\sup(S i)/(n)\sup(S i))) @@ -369,15 +431,13 @@ rewrite < eq_div_div_div_times |apply le_times_to_le_div2 [apply lt_O_exp.assumption |simplify. -*) -(* falso + theorem tech1: \forall a,b,n,m.O < m \to n/m \le b \to (a*n)/m \le a*b. intros. apply le_times_to_le_div2 [assumption | -*) theorem tech2: \forall n,m. O < n \to (exp (S n) m) / (exp n m) \le (n + m)/n. @@ -409,3 +469,4 @@ elim n |rewrite > true_to_sigma_p_Sn [apply (trans_le ? (m/n1+(log p (exp n1 m)))) [ +*) \ No newline at end of file