1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / Matita is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 include "datatypes/constructors.ma".
16 include "nat/minimization.ma".
17 include "nat/relevant_equations.ma".
18 include "nat/primes.ma".
19 include "nat/iteration2.ma".
20 include "nat/div_and_mod_diseq.ma".
22 definition log \def \lambda p,n.
23 max n (\lambda x.leb (exp p x) n).
25 theorem le_exp_log: \forall p,n. O < n \to
26 exp p (log p n) \le n.
30 apply (f_max_true (\lambda x.leb (exp p x) n)).
31 apply (ex_intro ? ? O).
34 |apply le_to_leb_true.simplify.assumption
38 theorem log_SO: \forall n. S O < n \to log n (S O) = O.
40 apply sym_eq.apply le_n_O_to_eq.
41 apply (le_exp_to_le n)
43 |simplify in ⊢ (? ? %).
49 theorem lt_to_log_O: \forall n,m. O < m \to m < n \to log n m = O.
51 apply sym_eq.apply le_n_O_to_eq.
53 apply (lt_exp_to_lt n)
54 [apply (le_to_lt_to_lt ? m);assumption
55 |simplify in ⊢ (? ? %).
57 apply (le_to_lt_to_lt ? m)
58 [apply le_exp_log.assumption
64 theorem lt_log_n_n: \forall p, n. S O < p \to O < n \to log p n < n.
67 [elim (le_to_or_lt_eq ? ? Hcut)
69 |absurd (exp p n \le n)
70 [rewrite < H2 in ⊢ (? (? ? %) ?).
78 |unfold log.apply le_max_n
82 theorem lt_O_log: \forall p,n. O < n \to p \le n \to O < log p n.
87 apply (leb_false_to_not_le ? ? ? H1).
88 rewrite > (exp_n_SO p).
89 apply (lt_max_to_false ? ? ? H2).
93 theorem le_log_n_n: \forall p,n. S O < p \to log p n \le n.
99 [assumption|apply lt_O_S]
103 theorem lt_exp_log: \forall p,n. S O < p \to n < exp p (S (log p n)).
105 [simplify.rewrite < times_n_SO.apply lt_to_le.assumption
107 apply leb_false_to_not_le.
108 apply (lt_max_to_false ? (S n1) (S (log p (S n1))))
109 [apply le_S_S.apply le_n
111 [assumption|apply lt_O_S]
116 theorem log_times1: \forall p,n,m. S O < p \to O < n \to O < m \to
117 log p (n*m) \le S(log p n+log p m).
119 unfold in ⊢ (? (% ? ?) ?).
120 apply f_false_to_le_max
121 [apply (ex_intro ? ? O).
124 |apply le_to_leb_true.
126 rewrite > times_n_SO.
127 apply le_times;assumption
130 apply lt_to_leb_false.
131 apply (lt_to_le_to_lt ? ((exp p (S(log p n)))*(exp p (S(log p m)))))
132 [apply lt_times;apply lt_exp_log;assumption
133 |rewrite < exp_plus_times.
135 [apply lt_to_le.assumption
144 theorem log_times: \forall p,n,m.S O < p \to log p (n*m) \le S(log p n+log p m).
149 [rewrite < times_n_O.
160 theorem log_times_l: \forall p,n,m.O < n \to O < m \to S O < p \to
161 log p n+log p m \le log p (n*m) .
163 unfold log in ⊢ (? ? (% ? ?)).
176 rewrite < times_n_SO.
181 |apply (trans_le ? (S n1 + S n2))
182 [apply le_plus;apply le_log_n_n;assumption
186 change in ⊢ (? % ?) with ((S n1)+n2).
189 change with (n1 < n1*S n2).
190 rewrite > times_n_SO in ⊢ (? % ?).
193 |apply le_S_S.assumption
198 |apply le_to_leb_true.
199 rewrite > exp_plus_times.
200 apply le_times;apply le_exp_log;assumption
204 theorem log_exp: \forall p,n,m.S O < p \to O < m \to
205 log p ((exp p n)*m)=n+log p m.
207 unfold log in ⊢ (? ? (% ? ?) ?).
208 apply max_spec_to_max.
218 rewrite > assoc_times.
219 apply (trans_le ? ((S(S O))*(p\sup n1*m)))
220 [apply le_S_times_SSO
221 [rewrite > (times_n_O O) in ⊢ (? % ?).
237 apply le_to_leb_true.
238 rewrite > exp_plus_times.
245 apply lt_to_leb_false.
246 apply (lt_to_le_to_lt ? ((exp p n)*(exp p (S(log p m)))))
248 [apply lt_O_exp.apply lt_to_le.assumption
249 |apply lt_exp_log.assumption
251 |rewrite < exp_plus_times.
253 [apply lt_to_le.assumption
254 |rewrite < plus_n_Sm.
261 theorem eq_log_exp: \forall p,n.S O < p \to
264 rewrite > times_n_SO in ⊢ (? ? (? ? %) ?).
267 [rewrite < plus_n_O.reflexivity
275 theorem log_exp1: \forall p,n,m.S O < p \to
276 log p (exp n m) \le m*S(log p n).
278 [simplify in ⊢ (? (? ? %) ?).
284 apply (trans_le ? (S (log p n+log p (n\sup n1))))
285 [apply log_times.assumption
293 theorem log_exp2: \forall p,n,m.S O < p \to O < n \to
294 m*(log p n) \le log p (exp n m).
297 apply (lt_exp_to_lt p)
299 |rewrite > sym_times.
300 rewrite < exp_exp_times.
301 apply (le_to_lt_to_lt ? (exp n m))
304 |simplify.apply le_times
316 lemma le_log_plus: \forall p,n.S O < p \to log p n \leq log p (S n).
317 intros;apply (bool_elim ? (leb (p*(exp p n)) (S n)))
318 [simplify;intro;rewrite > H1;simplify;apply (trans_le ? n)
319 [apply le_log_n_n;assumption
321 |intro;unfold log;simplify;rewrite > H1;simplify;apply le_max_f_max_g;
322 intros;apply le_to_leb_true;constructor 2;apply leb_true_to_le;assumption]
325 theorem le_log: \forall p,n,m. S O < p \to n \le m \to
329 |apply (trans_le ? ? ? H3);apply le_log_plus;assumption]
332 theorem log_div: \forall p,n,m. S O < p \to O < m \to m \le n \to
333 log p (n/m) \le log p n -log p m.
335 apply le_plus_to_minus_r.
336 apply (trans_le ? (log p ((n/m)*m)))
338 [apply le_times_to_le_div
340 |rewrite < times_n_SO.
348 |rewrite > (div_mod n m) in ⊢ (? ? %)
356 theorem log_n_n: \forall n. S O < n \to log n n = S O.
358 rewrite > exp_n_SO in ⊢ (? ? (? ? %) ?).
359 rewrite > times_n_SO in ⊢ (? ? (? ? %) ?).
370 theorem log_i_SSOn: \forall n,i. S O < n \to n < i \to i \le ((S(S O))*n) \to
371 log i ((S(S O))*n) = S O.
373 apply antisymmetric_le
374 [apply not_lt_to_le.intro.
375 apply (lt_to_not_le ((S(S O)) * n) (exp i (S(S O))))
378 [apply (le_to_lt_to_lt ? n);assumption
381 |apply (trans_le ? (exp i (log i ((S(S O))*n))))
383 [apply (ltn_to_ltO ? ? H1)
387 rewrite > (times_n_O O) in ⊢ (? % ?).
390 |apply lt_to_le.assumption
394 |apply (trans_le ? (log i i))
395 [rewrite < (log_n_n i) in ⊢ (? % ?)
397 [apply (trans_lt ? n);assumption
400 |apply (trans_lt ? n);assumption
403 [apply (trans_lt ? n);assumption
410 theorem exp_n_O: \forall n. O < n \to exp O n = O.
411 intros.apply (lt_O_n_elim ? H).intros.
412 simplify.reflexivity.
416 theorem tech1: \forall n,i.O < n \to
417 (exp (S n) (S(S i)))/(exp n (S i)) \le ((exp n i) + (exp (S n) (S i)))/(exp n i).
419 simplify in ⊢ (? (? ? %) ?).
420 rewrite < eq_div_div_div_times
422 [apply lt_O_exp.assumption
424 apply lt_times_to_lt_div.
425 change in ⊢ (? % ?) with ((exp (S n) (S i)) + n*(exp (S n) (S i))).
428 |apply (trans_le ? ((n)\sup(i)*(S n)\sup(S i)/(n)\sup(S i)))
429 [apply le_times_div_div_times.
430 apply lt_O_exp.assumption
431 |apply le_times_to_le_div2
432 [apply lt_O_exp.assumption
435 theorem tech1: \forall a,b,n,m.O < m \to
436 n/m \le b \to (a*n)/m \le a*b.
438 apply le_times_to_le_div2
442 theorem tech2: \forall n,m. O < n \to
443 (exp (S n) m) / (exp n m) \le (n + m)/n.
446 [rewrite < plus_n_O.simplify.
447 rewrite > div_n_n.apply le_n
448 |apply le_times_to_le_div
450 |apply (trans_le ? (n*(S n)\sup(S n1)/(n)\sup(S n1)))
451 [apply le_times_div_div_times.
453 |simplify in ⊢ (? (? ? %) ?).
454 rewrite > sym_times in ⊢ (? (? ? %) ?).
455 rewrite < eq_div_div_div_times
456 [apply le_times_to_le_div2
461 theorem le_log_sigma_p:\forall n,m,p. O < m \to S O < p \to
462 log p (exp n m) \le sigma_p n (\lambda i.true) (\lambda i. (m / i)).
469 |rewrite > true_to_sigma_p_Sn
470 [apply (trans_le ? (m/n1+(log p (exp n1 m))))