]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/nat/log.ma
Progress.
[helm.git] / helm / software / 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 set "baseuri" "cic:/matita/nat/log".
16
17 include "datatypes/constructors.ma".
18 include "nat/minimization.ma".
19 include "nat/relevant_equations.ma".
20 include "nat/primes.ma".
21 include "nat/iteration2.ma".
22 include "nat/div_and_mod_diseq.ma".
23
24 definition log \def \lambda p,n.
25 max n (\lambda x.leb (exp p x) n).
26
27 theorem le_exp_log: \forall p,n. O < n \to
28 exp p (log p n) \le n.
29 intros.
30 apply leb_true_to_le.
31 unfold log.
32 apply (f_max_true (\lambda x.leb (exp p x) n)).
33 apply (ex_intro ? ? O).
34 split
35   [apply le_O_n
36   |apply le_to_leb_true.simplify.assumption
37   ]
38 qed.
39
40 theorem log_SO: \forall n. S O < n \to log n (S O) = O.
41 intros.
42 apply sym_eq.apply le_n_O_to_eq.
43 apply (le_exp_to_le n)
44   [assumption
45   |simplify in ⊢ (? ? %).
46    apply le_exp_log.
47    apply le_n
48   ]
49 qed.
50
51 theorem lt_to_log_O: \forall n,m. O < m \to m < n \to log n m = O.
52 intros.
53 apply sym_eq.apply le_n_O_to_eq.
54 apply le_S_S_to_le.
55 apply (lt_exp_to_lt n)
56   [apply (le_to_lt_to_lt ? m);assumption
57   |simplify in ⊢ (? ? %).
58    rewrite < times_n_SO.
59    apply (le_to_lt_to_lt ? m)
60     [apply le_exp_log.assumption
61     |assumption
62     ]
63   ]
64 qed.
65
66 theorem lt_log_n_n: \forall p, n. S O < p \to O < n \to log p n < n.
67 intros.
68 cut (log p n \le n)
69   [elim (le_to_or_lt_eq ? ? Hcut)
70     [assumption
71     |absurd (exp p n \le n)
72       [rewrite < H2 in ⊢ (? (? ? %) ?).
73        apply le_exp_log.
74        assumption
75       |apply lt_to_not_le.
76        apply lt_m_exp_nm.
77        assumption
78       ]
79     ]
80   |unfold log.apply le_max_n
81   ]
82 qed.
83
84 theorem lt_O_log: \forall p,n. O < n \to p \le n \to O < log p n.
85 intros.
86 unfold log.
87 apply not_lt_to_le.
88 intro.
89 apply (leb_false_to_not_le ? ? ? H1).
90 rewrite > (exp_n_SO p).
91 apply (lt_max_to_false ? ? ? H2).
92 assumption.
93 qed.
94
95 theorem le_log_n_n: \forall p,n. S O < p \to log p n \le n.
96 intros.
97 cases n
98   [apply le_n
99   |apply lt_to_le.
100    apply lt_log_n_n
101     [assumption|apply lt_O_S]
102   ]
103 qed.
104
105 theorem lt_exp_log: \forall p,n. S O < p \to n < exp p (S (log p n)).
106 intros.cases n
107   [simplify.rewrite < times_n_SO.apply lt_to_le.assumption
108   |apply not_le_to_lt.
109    apply leb_false_to_not_le.
110    apply (lt_max_to_false ? (S n1) (S (log p (S n1))))
111     [apply le_S_S.apply le_n
112     |apply lt_log_n_n
113       [assumption|apply lt_O_S]
114     ]
115   ]
116 qed.
117
118 theorem log_times1: \forall p,n,m. S O < p \to O < n \to O < m \to
119 log p (n*m) \le S(log p n+log p m).
120 intros.
121 unfold in ⊢ (? (% ? ?) ?).
122 apply f_false_to_le_max
123   [apply (ex_intro ? ? O).
124    split
125     [apply le_O_n
126     |apply le_to_leb_true.
127      simplify.
128      rewrite > times_n_SO.
129      apply le_times;assumption
130     ]
131   |intros.
132    apply lt_to_leb_false.
133    apply (lt_to_le_to_lt ? ((exp p (S(log p n)))*(exp p (S(log p m)))))
134     [apply lt_times;apply lt_exp_log;assumption
135     |rewrite < exp_plus_times.
136      apply le_exp
137       [apply lt_to_le.assumption
138       |simplify.
139        rewrite < plus_n_Sm.
140        assumption
141       ]
142     ]
143   ]
144 qed.
145   
146 theorem log_times: \forall p,n,m.S O < p \to log p (n*m) \le S(log p n+log p m).
147 intros.
148 cases n
149   [apply le_O_n
150   |cases m
151     [rewrite < times_n_O.
152      apply le_O_n
153     |apply log_times1
154       [assumption
155       |apply lt_O_S
156       |apply lt_O_S
157       ]
158     ]
159   ]
160 qed.
161
162 theorem log_exp: \forall p,n,m.S O < p \to O < m \to
163 log p ((exp p n)*m)=n+log p m.
164 intros.
165 unfold log in ⊢ (? ? (% ? ?) ?).
166 apply max_spec_to_max.
167 unfold max_spec.
168 split
169   [split
170     [elim n
171       [simplify.
172        rewrite < plus_n_O.
173        apply le_log_n_n.
174        assumption
175       |simplify.
176        rewrite > assoc_times.
177        apply (trans_le ? ((S(S O))*(p\sup n1*m)))
178         [apply le_S_times_SSO
179           [rewrite > (times_n_O O) in ⊢ (? % ?).
180            apply lt_times
181             [apply lt_O_exp.
182              apply lt_to_le.
183              assumption
184             |assumption
185             ]
186           |assumption
187           ]
188         |apply le_times
189           [assumption
190           |apply le_n
191           ]
192         ]
193       ]
194     |simplify.
195      apply le_to_leb_true.
196      rewrite > exp_plus_times.
197      apply le_times_r.
198      apply le_exp_log.
199      assumption
200     ]
201   |intros.
202    simplify.
203    apply lt_to_leb_false.
204    apply (lt_to_le_to_lt ? ((exp p n)*(exp p (S(log p m)))))
205     [apply lt_times_r1
206       [apply lt_O_exp.apply lt_to_le.assumption
207       |apply lt_exp_log.assumption
208       ]
209     |rewrite < exp_plus_times.
210      apply le_exp
211       [apply lt_to_le.assumption
212       |rewrite < plus_n_Sm.
213        assumption
214       ]
215     ]
216   ]
217 qed.
218
219 theorem eq_log_exp: \forall p,n.S O < p \to
220 log p (exp p n)=n.
221 intros.
222 rewrite > times_n_SO in ⊢ (? ? (? ? %) ?).
223 rewrite > log_exp
224   [rewrite > log_SO
225     [rewrite < plus_n_O.reflexivity
226     |assumption
227     ]
228   |assumption
229   |apply le_n
230   ]
231 qed.
232
233 theorem log_exp1: \forall p,n,m.S O < p \to
234 log p (exp n m) \le m*S(log p n).
235 intros.elim m
236   [simplify in ⊢ (? (? ? %) ?).
237    rewrite > log_SO
238     [apply le_O_n
239     |assumption
240     ]
241   |simplify.
242    apply (trans_le ? (S (log p n+log p (n\sup n1))))
243     [apply log_times.assumption
244     |apply le_S_S.
245      apply le_plus_r.
246      assumption
247     ]
248   ]
249 qed.
250     
251 theorem log_exp2: \forall p,n,m.S O < p \to O < n \to
252 m*(log p n) \le log p (exp n m).
253 intros.
254 apply le_S_S_to_le.
255 apply (lt_exp_to_lt p)
256   [assumption
257   |rewrite > sym_times.
258    rewrite < exp_exp_times.
259    apply (le_to_lt_to_lt ? (exp n m))
260     [elim m
261       [simplify.apply le_n
262       |simplify.apply le_times
263         [apply le_exp_log.
264          assumption
265         |assumption
266         ]
267       ]
268     |apply lt_exp_log.
269      assumption
270     ]
271   ]
272 qed.
273
274 lemma le_log_plus: \forall p,n.S O < p \to log p n \leq log p (S n).
275 intros;apply (bool_elim ? (leb (p*(exp p n)) (S n)))
276   [simplify;intro;rewrite > H1;simplify;apply (trans_le ? n)
277     [apply le_log_n_n;assumption
278     |apply le_n_Sn]
279   |intro;unfold log;simplify;rewrite > H1;simplify;apply le_max_f_max_g;
280    intros;apply le_to_leb_true;constructor 2;apply leb_true_to_le;assumption]
281 qed.
282
283 theorem le_log: \forall p,n,m. S O < p \to n \le m \to 
284 log p n \le log p m.
285 intros.elim H1
286   [constructor 1
287   |apply (trans_le ? ? ? H3);apply le_log_plus;assumption]
288 qed.
289
290 theorem log_n_n: \forall n. S O < n \to log n n = S O.
291 intros.
292 rewrite > exp_n_SO in ⊢ (? ? (? ? %) ?).
293 rewrite > times_n_SO in ⊢ (? ? (? ? %) ?).
294 rewrite > log_exp
295   [rewrite > log_SO
296     [reflexivity
297     |assumption
298     ]
299   |assumption
300   |apply le_n
301   ]
302 qed.
303
304 theorem log_i_SSOn: \forall n,i. S O < n \to n < i \to i \le ((S(S O))*n) \to
305 log i ((S(S O))*n) = S O.
306 intros.
307 apply antisymmetric_le
308   [apply not_lt_to_le.intro.
309    apply (lt_to_not_le ((S(S O)) * n) (exp i (S(S O))))
310     [rewrite > exp_SSO.
311      apply lt_times
312       [apply (le_to_lt_to_lt ? n);assumption
313       |assumption
314       ]
315     |apply (trans_le ? (exp i (log i ((S(S O))*n))))
316       [apply le_exp
317         [apply (ltn_to_ltO ? ? H1)
318         |assumption
319         ]
320       |apply le_exp_log.
321        rewrite > (times_n_O O) in ⊢ (? % ?).
322        apply lt_times
323         [apply lt_O_S
324         |apply lt_to_le.assumption
325         ]
326       ]
327     ]
328   |apply (trans_le ? (log i i))
329     [rewrite < (log_n_n i) in ⊢ (? % ?)
330       [apply le_log
331         [apply (trans_lt ? n);assumption
332         |apply le_n
333         ]
334       |apply (trans_lt ? n);assumption
335       ]
336     |apply le_log
337       [apply (trans_lt ? n);assumption
338       |assumption
339       ]
340     ]
341   ]
342 qed.
343
344 theorem exp_n_O: \forall n. O < n \to exp O n = O.
345 intros.apply (lt_O_n_elim ? H).intros.
346 simplify.reflexivity.
347 qed.
348
349 (*
350 theorem tech1: \forall n,i.O < n \to
351 (exp (S n) (S(S i)))/(exp n (S i)) \le ((exp n i) + (exp (S n) (S i)))/(exp n i).
352 intros.
353 simplify in ⊢ (? (? ? %) ?).
354 rewrite < eq_div_div_div_times
355   [apply monotonic_div
356     [apply lt_O_exp.assumption
357     |apply le_S_S_to_le.
358      apply lt_times_to_lt_div.
359      change in ⊢ (? % ?) with ((exp (S n) (S i)) + n*(exp (S n) (S i))).
360       
361       
362   |apply (trans_le ? ((n)\sup(i)*(S n)\sup(S i)/(n)\sup(S i)))
363     [apply le_times_div_div_times.
364      apply lt_O_exp.assumption
365     |apply le_times_to_le_div2
366       [apply lt_O_exp.assumption
367       |simplify.
368
369 theorem tech1: \forall a,b,n,m.O < m \to
370 n/m \le b \to (a*n)/m \le a*b.
371 intros.
372 apply le_times_to_le_div2
373   [assumption
374   |
375
376 theorem tech2: \forall n,m. O < n \to 
377 (exp (S n) m) / (exp n m) \le (n + m)/n.
378 intros.
379 elim m
380   [rewrite < plus_n_O.simplify.
381    rewrite > div_n_n.apply le_n
382   |apply le_times_to_le_div
383     [assumption
384     |apply (trans_le ? (n*(S n)\sup(S n1)/(n)\sup(S n1)))
385       [apply le_times_div_div_times.
386        apply lt_O_exp
387       |simplify in ⊢ (? (? ? %) ?).
388        rewrite > sym_times in ⊢ (? (? ? %) ?). 
389        rewrite < eq_div_div_div_times
390         [apply le_times_to_le_div2
391           [assumption
392           |
393       
394       
395 theorem le_log_sigma_p:\forall n,m,p. O < m \to S O < p \to
396 log p (exp n m) \le sigma_p n (\lambda i.true) (\lambda i. (m / i)).
397 intros.
398 elim n
399   [rewrite > exp_n_O
400     [simplify.apply le_n
401     |assumption
402     ]
403   |rewrite > true_to_sigma_p_Sn
404     [apply (trans_le ? (m/n1+(log p (exp n1 m))))
405       [
406 *)