]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/log.ma
matita 0.5.1 tagged
[helm.git] / matita / library / nat / log.ma
1 (**************************************************************************)
2 (*       ___                                                                *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        Matita is distributed under the terms of the          *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
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".
21
22 definition log \def \lambda p,n.
23 max n (\lambda x.leb (exp p x) n).
24
25 theorem le_exp_log: \forall p,n. O < n \to
26 exp p (log p n) \le n.
27 intros.
28 apply leb_true_to_le.
29 unfold log.
30 apply (f_max_true (\lambda x.leb (exp p x) n)).
31 apply (ex_intro ? ? O).
32 split
33   [apply le_O_n
34   |apply le_to_leb_true.simplify.assumption
35   ]
36 qed.
37
38 theorem log_SO: \forall n. S O < n \to log n (S O) = O.
39 intros.
40 apply sym_eq.apply le_n_O_to_eq.
41 apply (le_exp_to_le n)
42   [assumption
43   |simplify in ⊢ (? ? %).
44    apply le_exp_log.
45    apply le_n
46   ]
47 qed.
48
49 theorem lt_to_log_O: \forall n,m. O < m \to m < n \to log n m = O.
50 intros.
51 apply sym_eq.apply le_n_O_to_eq.
52 apply le_S_S_to_le.
53 apply (lt_exp_to_lt n)
54   [apply (le_to_lt_to_lt ? m);assumption
55   |simplify in ⊢ (? ? %).
56    rewrite < times_n_SO.
57    apply (le_to_lt_to_lt ? m)
58     [apply le_exp_log.assumption
59     |assumption
60     ]
61   ]
62 qed.
63
64 theorem lt_log_n_n: \forall p, n. S O < p \to O < n \to log p n < n.
65 intros.
66 cut (log p n \le n)
67   [elim (le_to_or_lt_eq ? ? Hcut)
68     [assumption
69     |absurd (exp p n \le n)
70       [rewrite < H2 in ⊢ (? (? ? %) ?).
71        apply le_exp_log.
72        assumption
73       |apply lt_to_not_le.
74        apply lt_m_exp_nm.
75        assumption
76       ]
77     ]
78   |unfold log.apply le_max_n
79   ]
80 qed.
81
82 theorem lt_O_log: \forall p,n. O < n \to p \le n \to O < log p n.
83 intros.
84 unfold log.
85 apply not_lt_to_le.
86 intro.
87 apply (leb_false_to_not_le ? ? ? H1).
88 rewrite > (exp_n_SO p).
89 apply (lt_max_to_false ? ? ? H2).
90 assumption.
91 qed.
92
93 theorem le_log_n_n: \forall p,n. S O < p \to log p n \le n.
94 intros.
95 cases n
96   [apply le_n
97   |apply lt_to_le.
98    apply lt_log_n_n
99     [assumption|apply lt_O_S]
100   ]
101 qed.
102
103 theorem lt_exp_log: \forall p,n. S O < p \to n < exp p (S (log p n)).
104 intros.cases n
105   [simplify.rewrite < times_n_SO.apply lt_to_le.assumption
106   |apply not_le_to_lt.
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
110     |apply lt_log_n_n
111       [assumption|apply lt_O_S]
112     ]
113   ]
114 qed.
115
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).
118 intros.
119 unfold in ⊢ (? (% ? ?) ?).
120 apply f_false_to_le_max
121   [apply (ex_intro ? ? O).
122    split
123     [apply le_O_n
124     |apply le_to_leb_true.
125      simplify.
126      rewrite > times_n_SO.
127      apply le_times;assumption
128     ]
129   |intros.
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.
134      apply le_exp
135       [apply lt_to_le.assumption
136       |simplify.
137        rewrite < plus_n_Sm.
138        assumption
139       ]
140     ]
141   ]
142 qed.
143   
144 theorem log_times: \forall p,n,m.S O < p \to log p (n*m) \le S(log p n+log p m).
145 intros.
146 cases n
147   [apply le_O_n
148   |cases m
149     [rewrite < times_n_O.
150      apply le_O_n
151     |apply log_times1
152       [assumption
153       |apply lt_O_S
154       |apply lt_O_S
155       ]
156     ]
157   ]
158 qed.
159
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) .
162 intros.
163 unfold log in ⊢ (? ? (% ? ?)).
164 apply f_m_to_le_max
165   [elim H
166     [rewrite > log_SO
167       [simplify.
168        rewrite < plus_n_O.
169        apply le_log_n_n.
170        assumption
171       |assumption
172       ]
173     |elim H1
174       [rewrite > log_SO
175         [rewrite < plus_n_O.
176          rewrite < times_n_SO.
177          apply le_log_n_n.
178          assumption
179         |assumption
180         ]
181       |apply (trans_le ? (S n1 + S n2))
182         [apply le_plus;apply le_log_n_n;assumption
183         |simplify.
184          apply le_S_S.
185          rewrite < plus_n_Sm.
186          change in ⊢ (? % ?) with ((S n1)+n2).
187          rewrite > sym_plus.
188          apply le_plus_r.
189          change with (n1 < n1*S n2).
190          rewrite > times_n_SO in ⊢ (? % ?).
191          apply lt_times_r1
192           [assumption
193           |apply le_S_S.assumption
194           ]
195         ]
196       ]
197     ]
198   |apply le_to_leb_true.
199    rewrite > exp_plus_times.
200    apply le_times;apply le_exp_log;assumption
201   ]
202 qed.
203
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.
206 intros.
207 unfold log in ⊢ (? ? (% ? ?) ?).
208 apply max_spec_to_max.
209 unfold max_spec.
210 split
211   [split
212     [elim n
213       [simplify.
214        rewrite < plus_n_O.
215        apply le_log_n_n.
216        assumption
217       |simplify.
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 ⊢ (? % ?).
222            apply lt_times
223             [apply lt_O_exp.
224              apply lt_to_le.
225              assumption
226             |assumption
227             ]
228           |assumption
229           ]
230         |apply le_times
231           [assumption
232           |apply le_n
233           ]
234         ]
235       ]
236     |simplify.
237      apply le_to_leb_true.
238      rewrite > exp_plus_times.
239      apply le_times_r.
240      apply le_exp_log.
241      assumption
242     ]
243   |intros.
244    simplify.
245    apply lt_to_leb_false.
246    apply (lt_to_le_to_lt ? ((exp p n)*(exp p (S(log p m)))))
247     [apply lt_times_r1
248       [apply lt_O_exp.apply lt_to_le.assumption
249       |apply lt_exp_log.assumption
250       ]
251     |rewrite < exp_plus_times.
252      apply le_exp
253       [apply lt_to_le.assumption
254       |rewrite < plus_n_Sm.
255        assumption
256       ]
257     ]
258   ]
259 qed.
260
261 theorem eq_log_exp: \forall p,n.S O < p \to
262 log p (exp p n)=n.
263 intros.
264 rewrite > times_n_SO in ⊢ (? ? (? ? %) ?).
265 rewrite > log_exp
266   [rewrite > log_SO
267     [rewrite < plus_n_O.reflexivity
268     |assumption
269     ]
270   |assumption
271   |apply le_n
272   ]
273 qed.
274
275 theorem log_exp1: \forall p,n,m.S O < p \to
276 log p (exp n m) \le m*S(log p n).
277 intros.elim m
278   [simplify in ⊢ (? (? ? %) ?).
279    rewrite > log_SO
280     [apply le_O_n
281     |assumption
282     ]
283   |simplify.
284    apply (trans_le ? (S (log p n+log p (n\sup n1))))
285     [apply log_times.assumption
286     |apply le_S_S.
287      apply le_plus_r.
288      assumption
289     ]
290   ]
291 qed.
292     
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).
295 intros.
296 apply le_S_S_to_le.
297 apply (lt_exp_to_lt p)
298   [assumption
299   |rewrite > sym_times.
300    rewrite < exp_exp_times.
301    apply (le_to_lt_to_lt ? (exp n m))
302     [elim m
303       [simplify.apply le_n
304       |simplify.apply le_times
305         [apply le_exp_log.
306          assumption
307         |assumption
308         ]
309       ]
310     |apply lt_exp_log.
311      assumption
312     ]
313   ]
314 qed.
315
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
320     |apply le_n_Sn]
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]
323 qed.
324
325 theorem le_log: \forall p,n,m. S O < p \to n \le m \to 
326 log p n \le log p m.
327 intros.elim H1
328   [constructor 1
329   |apply (trans_le ? ? ? H3);apply le_log_plus;assumption]
330 qed.
331          
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.
334 intros.
335 apply le_plus_to_minus_r.
336 apply (trans_le ? (log p ((n/m)*m)))
337   [apply log_times_l
338     [apply le_times_to_le_div
339       [assumption
340       |rewrite < times_n_SO.
341        assumption
342       ]
343     |assumption
344     |assumption
345     ]
346   |apply le_log
347     [assumption
348     |rewrite > (div_mod n m) in ⊢ (? ? %)
349       [apply le_plus_n_r
350       |assumption
351       ]
352     ]
353   ]
354 qed.
355
356 theorem log_n_n: \forall n. S O < n \to log n n = S O.
357 intros.
358 rewrite > exp_n_SO in ⊢ (? ? (? ? %) ?).
359 rewrite > times_n_SO in ⊢ (? ? (? ? %) ?).
360 rewrite > log_exp
361   [rewrite > log_SO
362     [reflexivity
363     |assumption
364     ]
365   |assumption
366   |apply le_n
367   ]
368 qed.
369
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.
372 intros.
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))))
376     [rewrite > exp_SSO.
377      apply lt_times
378       [apply (le_to_lt_to_lt ? n);assumption
379       |assumption
380       ]
381     |apply (trans_le ? (exp i (log i ((S(S O))*n))))
382       [apply le_exp
383         [apply (ltn_to_ltO ? ? H1)
384         |assumption
385         ]
386       |apply le_exp_log.
387        rewrite > (times_n_O O) in ⊢ (? % ?).
388        apply lt_times
389         [apply lt_O_S
390         |apply lt_to_le.assumption
391         ]
392       ]
393     ]
394   |apply (trans_le ? (log i i))
395     [rewrite < (log_n_n i) in ⊢ (? % ?)
396       [apply le_log
397         [apply (trans_lt ? n);assumption
398         |apply le_n
399         ]
400       |apply (trans_lt ? n);assumption
401       ]
402     |apply le_log
403       [apply (trans_lt ? n);assumption
404       |assumption
405       ]
406     ]
407   ]
408 qed.
409
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.
413 qed.
414
415 (*
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).
418 intros.
419 simplify in ⊢ (? (? ? %) ?).
420 rewrite < eq_div_div_div_times
421   [apply monotonic_div
422     [apply lt_O_exp.assumption
423     |apply le_S_S_to_le.
424      apply lt_times_to_lt_div.
425      change in ⊢ (? % ?) with ((exp (S n) (S i)) + n*(exp (S n) (S i))).
426       
427       
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
433       |simplify.
434
435 theorem tech1: \forall a,b,n,m.O < m \to
436 n/m \le b \to (a*n)/m \le a*b.
437 intros.
438 apply le_times_to_le_div2
439   [assumption
440   |
441
442 theorem tech2: \forall n,m. O < n \to 
443 (exp (S n) m) / (exp n m) \le (n + m)/n.
444 intros.
445 elim m
446   [rewrite < plus_n_O.simplify.
447    rewrite > div_n_n.apply le_n
448   |apply le_times_to_le_div
449     [assumption
450     |apply (trans_le ? (n*(S n)\sup(S n1)/(n)\sup(S n1)))
451       [apply le_times_div_div_times.
452        apply lt_O_exp
453       |simplify in ⊢ (? (? ? %) ?).
454        rewrite > sym_times in ⊢ (? (? ? %) ?). 
455        rewrite < eq_div_div_div_times
456         [apply le_times_to_le_div2
457           [assumption
458           |
459       
460       
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)).
463 intros.
464 elim n
465   [rewrite > exp_n_O
466     [simplify.apply le_n
467     |assumption
468     ]
469   |rewrite > true_to_sigma_p_Sn
470     [apply (trans_le ? (m/n1+(log p (exp n1 m))))
471       [
472 *)