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 "nat/sqrt.ma".
16 include "nat/chebyshev_teta.ma".
17 include "nat/chebyshev.ma".
19 include "list/sort.ma".
21 include "nat/sieve.ma".
23 let rec list_divides l n \def
26 | cons (m:nat) (tl:list nat) ⇒ orb (divides_b m n) (list_divides tl n) ].
28 definition lprim : nat \to list nat \def
29 \lambda n.let rec aux m acc \def
32 | S m1 => match (list_divides acc (n-m1)) with
34 | false => aux m1 (n-m1::acc)]]
37 let rec checker l \def
40 | cons h1 t1 => match t1 with
42 | cons h2 t2 => (andb (checker t1) (leb h1 (2*h2))) ]].
44 lemma checker_cons : \forall t,l.checker (t::l) = true \to checker l = true.
45 intros 2;simplify;intro;elim l in H ⊢ %
47 |change in H1 with (andb (checker (a::l1)) (leb t (a+(a+O))) = true);
48 apply (andb_true_true ? ? H1)]
51 theorem checker_sound : \forall l1,l2,l,x,y.l = l1@(x::y::l2) \to
52 checker l = true \to x \leq 2*y.
54 [simplify;intros 5;rewrite > H;simplify;intro;
55 apply leb_true_to_le;apply (andb_true_true_r ? ? H1);
56 |simplify;intros;rewrite > H1 in H2;lapply (checker_cons ? ? H2);
57 apply (H l2 ? ? ? ? Hletin);reflexivity]
60 definition bertrand \def \lambda n.
61 \exists p.n < p \land p \le 2*n \land (prime p).
63 definition not_bertrand \def \lambda n.
64 \forall p.n < p \to p \le 2*n \to \not (prime p).
67 lemma list_of_primes_SO: \forall l.list_of_primes 1 l \to
71 |apply False_ind.unfold in H.
72 absurd ((prime n) \land n \le 1)
77 apply (lt_to_not_le ? ? H4 H3)
83 lemma min_prim : \forall n.\exists p. n < p \land prime p \land
84 \forall q.prime q \to q < p \to q \leq n.
85 intro;elim (le_to_or_lt_eq ? ? (le_O_n n))
86 [apply (ex_intro ? ? (min_aux (S (n!)) (S n) primeb));
90 |apply primeb_true_to_prime;apply f_min_aux_true;elim (ex_prime n);
91 [apply (ex_intro ? ? a);elim H1;elim H2;split
94 |rewrite > plus_n_O;apply le_plus
97 |apply prime_to_primeb_true;assumption]
99 |intros;apply not_lt_to_le;intro;lapply (lt_min_aux_to_false ? ? ? ? H3 H2);
100 rewrite > (prime_to_primeb_true ? H1) in Hletin;destruct Hletin]
101 |apply (ex_intro ? ? 2);split
103 [rewrite < H;apply lt_O_S
104 |apply primeb_true_to_prime;reflexivity]
105 |intros;elim (lt_to_not_le ? ? H2);apply prime_to_lt_SO;assumption]]
108 theorem list_of_primes_to_bertrand: \forall n,pn,l.0 < n \to prime pn \to n <pn \to
109 list_of_primes pn l \to
110 (\forall p. prime p \to p \le pn \to in_list nat p l) \to
111 (\forall p. in_list nat p l \to 2 < p \to
112 \exists pp. in_list nat pp l \land pp < p \land p \le 2*pp) \to bertrand n.
115 apply (ex_intro ? ? a).
116 elim H6.clear H6.elim H7.clear H7.
120 |elim (le_to_or_lt_eq ? ? (prime_to_lt_SO ? H9))
122 [elim H10.clear H10.elim H11.clear H11.
123 apply (trans_le ? ? ? H12).
133 |apply not_lt_to_le.intro.
134 apply (lt_to_not_le ? ? H2).
140 apply O_lt_const_to_le_times_const.
148 let rec check_list l \def
150 [ nil \Rightarrow true
151 | cons (hd:nat) tl \Rightarrow
153 [ nil \Rightarrow eqb hd 2
154 | cons hd1 tl1 \Rightarrow
155 (leb (S hd1) hd \land leb hd (2*hd1) \land check_list tl)
160 lemma check_list1: \forall n,m,l.(check_list (n::m::l)) = true \to
161 m < n \land n \le 2*m \land (check_list (m::l)) = true \land ((check_list l) = true).
163 change in ⊢ (? ? % ?→?) with (leb (S m) n \land leb n (2*m) \land check_list (m::l)).
165 lapply (andb_true_true ? ? H) as H1.
166 lapply (andb_true_true_r ? ? H) as H2.clear H.
167 lapply (andb_true_true ? ? H1) as H3.
168 lapply (andb_true_true_r ? ? H1) as H4.clear H1.
172 [apply leb_true_to_le.assumption
173 |apply leb_true_to_le.assumption
177 |generalize in match H2.
180 |change in ⊢ (? ? % ?→?) with (leb (S n1) m \land leb m (2*n1) \land check_list (n1::l1)).
182 lapply (andb_true_true_r ? ? H) as H2.
188 theorem check_list2: \forall l. check_list l = true \to
189 \forall p. in_list nat p l \to 2 < p \to
190 \exists pp. in_list nat pp l \land pp < p \land p \le 2*pp.
192 [intros.apply False_ind.apply (not_in_list_nil ? ? H1)
194 [lapply (in_list_singleton_to_eq ? ? ? H2) as H4.
196 apply (lt_to_not_eq ? ? H3).
197 apply sym_eq.apply eqb_true_to_eq.
198 rewrite > H4.apply H1
199 |elim (check_list1 ? ? ? H1).clear H1.
202 elim (in_list_cons_case ? ? ? ? H2)
203 [apply (ex_intro ? ? n).
206 [apply in_list_cons.apply in_list_head
207 |rewrite > H1.assumption
209 |rewrite > H1.assumption
211 |elim (H H6 p H1 H3).clear H.
212 apply (ex_intro ? ? a1).
217 [apply in_list_cons.assumption
227 (* qualcosa che non va con gli S *)
228 lemma le_to_bertrand : \forall n.O < n \to n \leq exp 2 8 \to bertrand n.
230 apply (list_of_primes_to_bertrand ? (S(exp 2 8)) (sieve (S(exp 2 8))))
232 |apply primeb_true_to_prime.reflexivity
233 |apply (le_to_lt_to_lt ? ? ? H1).
235 |lapply (sieve_sound1 (S(exp 2 8))) as H
237 |apply leb_true_to_le.reflexivity
239 |intros.apply (sieve_sound2 ? ? H3 H2)
245 (*lemma pippo : \forall k,n.in_list ? (nth_prime (S k)) (sieve n) \to
246 \exists l.sieve n = l@((nth_prime (S k))::(sieve (nth_prime k))).
247 intros;elim H;elim H1;clear H H1;apply (ex_intro ? ? a);
248 cut (a1 = sieve (nth_prime k))
249 [rewrite < Hcut;assumption
250 |lapply (sieve_sorted n);generalize in match H2*)
252 (* old proof by Wilmer
253 lemma le_to_bertrand : \forall n.O < n \to n \leq exp 2 8 \to bertrand n.
255 elim (min_prim n);apply (ex_intro ? ? a);elim H2;elim H3;clear H2 H3;
257 [|apply not_lt_to_le;intro;apply (le_to_not_lt ? ? H1);apply (H4 ? ? H2);
258 apply primeb_true_to_prime;reflexivity]
262 |elim (prime_to_nth_prime a H6);generalize in match H2;cases a1
263 [simplify;intro;rewrite < H3;rewrite < plus_n_O;
264 change in \vdash (? % ?) with (1+1);apply le_plus;assumption
265 |intro;lapply (H4 (nth_prime n1))
266 [apply (trans_le ? (2*(nth_prime n1)))
268 cut (\exists l1,l2.sieve 257 = l1@((nth_prime (S n1))::((nth_prime n1)::l2)))
269 [elim Hcut1;elim H7;clear Hcut1 H7;
270 apply (checker_sound a2 a3 (sieve 257))
273 |elim (sieve_sound2 257 (nth_prime (S n1)) ? ?)
274 [elim (sieve_sound2 257 (nth_prime n1) ? ?)
276 cut (\forall p.in_list ? p (a3@(nth_prime n1::a4)) \to prime p)
277 [|rewrite < H9;intros;apply (in_list_sieve_to_prime 257 p ? H10);
278 apply leb_true_to_le;reflexivity]
279 apply (ex_intro ? ? a2);apply (ex_intro ? ? a4);
281 cut ((nth_prime n1)::a4 = a5)
282 [|generalize in match H10;
283 lapply (sieve_sorted 257);
284 generalize in match Hletin1;
285 rewrite > H9 in ⊢ (? %→? ? % ?→?);
286 generalize in match Hcut1;
287 generalize in match a2;
290 [change in H11 with (nth_prime n1::a4 = nth_prime (S n1)::a5);
291 destruct H11;elim (eq_to_not_lt ? ? Hcut2);
292 apply increasing_nth_prime
293 |change in H12 with (nth_prime n1::a4 = t::(l1@(nth_prime (S n1)::a5)));
295 change in H11 with (sorted_gt (nth_prime n1::l1@(nth_prime (S n1)::a5)));
296 lapply (sorted_to_minimum ? ? ? H11 (nth_prime (S n1)))
297 [unfold in Hletin2;elim (le_to_not_lt ? ? (lt_to_le ? ? Hletin2));
298 apply increasing_nth_prime
299 |apply (ex_intro ? ? l1);apply (ex_intro ? ? a5);reflexivity]]
301 [change in H12 with (t::(l@(nth_prime n1::a4)) = nth_prime (S n1)::a5);
302 destruct H12;cut (l = [])
303 [rewrite > Hcut2;reflexivity
304 |change in H11 with (sorted_gt (nth_prime (S n1)::(l@(nth_prime n1::a4))));
305 generalize in match H11;generalize in match H8;cases l;intros
307 |lapply (sorted_cons_to_sorted ? ? ? H13);
308 lapply (sorted_to_minimum ? ? ? H13 n2)
309 [simplify in Hletin2;lapply (sorted_to_minimum ? ? ? Hletin2 (nth_prime n1))
310 [unfold in Hletin3;unfold in Hletin4;
311 elim (lt_nth_prime_to_not_prime ? ? Hletin4 Hletin3);
313 apply (ex_intro ? ? [nth_prime (S n1)]);
314 apply (ex_intro ? ? (l2@(nth_prime n1::a4)));
316 |apply (ex_intro ? ? l2);apply (ex_intro ? ? a4);reflexivity]
317 |simplify;apply in_list_head]]]
318 |change in H13 with (t::(l@(nth_prime n1::a4)) = t1::(l2@(nth_prime (S n1)::a5)));
319 destruct H13;apply (H7 l2 ? ? Hcut3)
320 [intros;apply H8;simplify;apply in_list_cons;
323 apply (sorted_cons_to_sorted ? ? ? H12)]]]]
324 rewrite > Hcut2 in ⊢ (? ? ? (? ? ? (? ? ? %)));
326 |apply (trans_le ? ? ? Hletin);apply lt_to_le;
327 apply (trans_le ? ? ? H5 Hcut)
328 |apply prime_nth_prime]
329 |rewrite > H3;assumption
330 |apply prime_nth_prime]]
331 |apply le_times_r;assumption]
332 |apply prime_nth_prime
333 |rewrite < H3;apply increasing_nth_prime]]]
337 lemma not_not_bertrand_to_bertrand1: \forall n.
338 \lnot (not_bertrand n) \to \forall x. n \le x \to x \le 2*n \to
339 (\forall p.x < p \to p \le 2*n \to \not (prime p))
340 \to \exists p.n < p \land p \le x \land (prime p).
342 [apply False_ind.apply H.assumption
343 |apply (bool_elim ? (primeb (S n1)));intro
344 [apply (ex_intro ? ? (S n1)).
347 [apply le_S_S.assumption
350 |apply primeb_true_to_prime.assumption
355 apply (ex_intro ? ? a).
359 |apply le_S.assumption
363 |apply lt_to_le.assumption
364 |elim (le_to_or_lt_eq ? ? H7)
367 apply primeb_false_to_not_prime.
375 theorem not_not_bertrand_to_bertrand: \forall n.
376 \lnot (not_bertrand n) \to bertrand n.
377 unfold bertrand.intros.
378 apply (not_not_bertrand_to_bertrand1 ? ? (2*n))
380 |apply le_times_n.apply le_n_Sn
382 |intros.apply False_ind.
383 apply (lt_to_not_le ? ? H1 H2)
388 theorem divides_pi_p_to_divides: \forall p,n,b,g.prime p \to
389 divides p (pi_p n b g) \to \exists i. (i < n \and (b i = true \and
394 apply (le_to_not_lt p 1)
401 |apply (bool_elim ? (b n1));intro
402 [rewrite > (true_to_pi_p_Sn ? ? ? H3) in H2.
403 elim (divides_times_to_divides ? ? ? H1 H2)
404 [apply (ex_intro ? ? n1).
411 apply (ex_intro ? ? a).
413 [apply lt_to_le.apply le_S_S.assumption
417 |rewrite > (false_to_pi_p_Sn ? ? ? H3) in H2.
420 apply (ex_intro ? ? a).
422 [apply lt_to_le.apply le_S_S.assumption
429 theorem divides_B: \forall n,p.prime p \to p \divides (B n) \to
430 p \le n \land \exists i.mod (n /(exp p (S i))) 2 \neq O.
433 elim (divides_pi_p_to_divides ? ? ? ? H H1).
436 elim (divides_pi_p_to_divides ? ? ? ? H H5).clear H5.
441 [rewrite > Hcut.apply le_S_S_to_le.assumption
442 |apply (ex_intro ? ? a1).
445 change in H7:(? ? %) with (exp a ((n/(exp a (S a1))) \mod 2)).
449 [apply divides_to_le[apply lt_O_S|assumption]
450 |apply lt_to_not_le.elim H.assumption
453 |apply (divides_exp_to_eq ? ? ? H ? H7).
454 apply primeb_true_to_prime.
460 definition k \def \lambda n,p.
461 sigma_p (log p n) (λi:nat.true) (λi:nat.((n/(exp p (S i))\mod 2))).
463 theorem le_k: \forall n,p. k n p \le log p n.
464 intros.unfold k.elim (log p n)
466 |rewrite > true_to_sigma_p_Sn
467 [rewrite > plus_n_SO.
468 rewrite > sym_plus in ⊢ (? ? %).
481 \lambda n. pi_p (S n) primeb (\lambda p.(exp p (k n p))).
483 theorem eq_B_B1: \forall n. B n = B1 n.
484 intros.unfold B.unfold B1.
492 definition B_split1 \def \lambda n.
493 pi_p (S n) primeb (\lambda p.(exp p (bool_to_nat (leb (k n p) 1)* (k n p)))).
495 definition B_split2 \def \lambda n.
496 pi_p (S n) primeb (\lambda p.(exp p (bool_to_nat (leb 2 (k n p))* (k n p)))).
498 theorem eq_B1_times_B_split1_B_split2: \forall n.
499 B1 n = B_split1 n * B_split2 n.
500 intro.unfold B1.unfold B_split1.unfold B_split2.
501 rewrite < times_pi_p.
504 |intros.apply (bool_elim ? (leb (k n x) 1));intro
505 [rewrite > (lt_to_leb_false 2 (k n x))
506 [simplify.rewrite < plus_n_O.
507 rewrite < times_n_SO.reflexivity
508 |apply le_S_S.apply leb_true_to_le.assumption
510 |rewrite > (le_to_leb_true 2 (k n x))
511 [simplify.rewrite < plus_n_O.
512 rewrite < plus_n_O.reflexivity
513 |apply not_le_to_lt.apply leb_false_to_not_le.assumption
519 lemma lt_div_to_times: \forall n,m,q. O < q \to n/q < m \to n < q*m.
523 intro.apply (lt_to_not_le ? ? H1).
524 apply le_times_to_le_div;assumption
525 |apply (ltn_to_ltO ? ? H1)
529 lemma lt_to_div_O:\forall n,m. n < m \to n / m = O.
531 apply (div_mod_spec_to_eq n m (n/m) (n \mod m) O n)
532 [apply div_mod_spec_div_mod.
533 apply (ltn_to_ltO ? ? H)
534 |apply div_mod_spec_intro
541 (* the value of n could be smaller *)
542 lemma k1: \forall n,p. 18 \le n \to p \le n \to 2*n/ 3 < p\to k (2*n) p = O.
546 |rewrite > true_to_sigma_p_Sn
551 cut (2*n/p = 2) as H4
552 [rewrite > H4.reflexivity
553 |apply lt_to_le_times_to_lt_S_to_div
554 [apply (ltn_to_ltO ? ? H2)
555 |rewrite < sym_times.
558 |rewrite > sym_times in ⊢ (? ? %).
559 apply lt_div_to_times
565 |cut (2*n/(p)\sup(S (S n2)) = O) as H4
566 [rewrite > H4.reflexivity
568 apply (le_to_lt_to_lt ? (exp ((2*n)/3) 2))
569 [apply (le_times_to_le (exp 3 2))
570 [apply leb_true_to_le.reflexivity
571 |rewrite > sym_times in ⊢ (? ? %).
573 apply (trans_le ? (exp n 2))
574 [rewrite < assoc_times.
575 rewrite > exp_SSO in ⊢ (? ? %).
578 |apply monotonic_exp1.
579 apply (le_plus_to_le 3).
580 change in ⊢ (? ? %) with ((S(2*n/3))*3).
581 apply (trans_le ? (2*n))
582 [simplify in ⊢ (? ? %).
585 apply (trans_le ? 18 ? ? H).
586 apply leb_true_to_le.reflexivity
593 |apply (lt_to_le_to_lt ? (exp p 2))
599 [apply (ltn_to_ltO ? ? H2)
600 |apply le_S_S.apply le_S_S.apply le_O_n
611 theorem le_B_split1_teta:\forall n.18 \le n \to not_bertrand n \to
612 B_split1 (2*n) \le teta (2 * n / 3).
613 intros. unfold B_split1.unfold teta.
614 apply (trans_le ? (pi_p (S (2*n)) primeb (λp:nat.(p)\sup(bool_to_nat (eqb (k (2*n) p) 1)))))
615 [apply le_pi_p.intros.
617 [apply prime_to_lt_O.apply primeb_true_to_prime.assumption
618 |apply (bool_elim ? (leb (k (2*n) i) 1));intro
619 [elim (le_to_or_lt_eq ? ? (leb_true_to_le ? ? H4))
620 [lapply (le_S_S_to_le ? ? H5) as H6.
621 apply (le_n_O_elim ? H6).
624 |rewrite > (eq_to_eqb_true ? ? H5).
625 rewrite > H5.apply le_n
630 |apply (trans_le ? (pi_p (S (2*n/3)) primeb (λp:nat.(p)\sup(bool_to_nat (eqb (k (2*n) p) 1)))))
631 [apply (eq_ind ? ? ? (le_n ?)).
632 apply or_false_eq_SO_to_eq_pi_p
634 apply le_times_to_le_div2
636 |rewrite > sym_times in ⊢ (? ? %).
638 apply leb_true_to_le.reflexivity
641 unfold not_bertrand in H1.
642 elim (decidable_le (S n) i)
644 apply not_prime_to_primeb_false.
647 |apply le_S_S_to_le.assumption
654 apply not_le_to_lt.assumption
659 |apply le_pi_p.intros.
660 elim (eqb (k (2*n) i) 1)
661 [rewrite < exp_n_SO.apply le_n
662 |simplify.apply prime_to_lt_O.
663 apply primeb_true_to_prime.
670 theorem le_B_split2_exp: \forall n. exp 2 7 \le n \to
671 B_split2 (2*n) \le exp (2*n) (pred(sqrt(2*n)/2)).
672 intros.unfold B_split2.
674 [apply (trans_le ? (pi_p (S (sqrt (2*n))) primeb
675 (λp:nat.(p)\sup(bool_to_nat (leb 2 (k (2*n) p))*k (2*n) p))))
676 [apply (eq_ind ? ? ? (le_n ?)).
677 apply or_false_eq_SO_to_eq_pi_p
681 apply (bool_elim ? (leb 2 (k (2*n) i)));intro
683 apply (lt_to_not_le ? ? H1).unfold sqrt.
685 [apply le_S_S_to_le.assumption
686 |apply le_to_leb_true.
688 apply not_lt_to_le.intro.
689 apply (le_to_not_lt 2 (log i (2*n)))
690 [apply (trans_le ? (k (2*n) i))
691 [apply leb_true_to_le.assumption
694 |apply le_S_S.unfold log.apply f_false_to_le_max
695 [apply (ex_intro ? ? O).split
697 |apply le_to_leb_true.simplify.
703 |intros.apply lt_to_leb_false.
704 apply (lt_to_le_to_lt ? (exp i 2))
707 [apply (ltn_to_ltO ? ? H1)
717 |apply (trans_le ? (pi_p (S (sqrt (2*n))) primeb (λp:nat.2*n)))
718 [apply le_pi_p.intros.
719 apply (trans_le ? (exp i (log i (2*n))))
721 [apply prime_to_lt_O.
722 apply primeb_true_to_prime.
724 |apply (bool_elim ? (leb 2 (k (2*n) i)));intro
725 [simplify in ⊢ (? (? % ?) ?).
727 rewrite < times_n_SO.
733 rewrite > (times_n_O O) in ⊢ (? % ?).
739 |apply (trans_le ? (exp (2*n) (prim(sqrt (2*n)))))
741 apply (eq_ind ? ? ? (le_n ?)).
744 [rewrite > (times_n_O O) in ⊢ (? % ?).
752 [apply (trans_le ? (2*(exp 2 7)))
753 [apply leb_true_to_le.reflexivity
754 |apply le_times_r.assumption
756 |apply le_to_leb_true.
757 apply (trans_le ? (2*(exp 2 7)))
758 [apply leb_true_to_le.reflexivity
759 |apply le_times_r.assumption
766 |apply (lt_to_le_to_lt ? ? ? ? H).
767 apply leb_true_to_le.reflexivity
771 theorem not_bertrand_to_le_B:
772 \forall n.exp 2 7 \le n \to not_bertrand n \to
773 B (2*n) \le (exp 2 (2*(2 * n / 3)))*(exp (2*n) (pred(sqrt(2*n)/2))).
776 rewrite > eq_B1_times_B_split1_B_split2.
778 [apply (trans_le ? (teta ((2*n)/3)))
779 [apply le_B_split1_teta
780 [apply (trans_le ? ? ? ? H).
781 apply leb_true_to_le.reflexivity
786 |apply le_B_split2_exp.
792 theorem not_bertrand_to_le1:
793 \forall n.18 \le n \to not_bertrand n \to
794 exp 2 (2*n) \le (exp 2 (2*(2 * n / 3)))*(exp (2*n) (S(sqrt(2*n)))).
797 theorem le_times_div_m_m: \forall n,m. O < m \to n/m*m \le n.
799 rewrite > (div_mod n m) in ⊢ (? ? %)
805 theorem not_bertrand_to_le1:
806 \forall n.exp 2 7 \le n \to not_bertrand n \to
807 (exp 2 (2*n / 3)) \le (exp (2*n) (sqrt(2*n)/2)).
809 apply (le_times_to_le (exp 2 (2*(2 * n / 3))))
810 [apply lt_O_exp.apply lt_O_S
811 |rewrite < exp_plus_times.
812 apply (trans_le ? (exp 2 (2*n)))
816 change in ⊢ (? % ?) with (3*(2*n/3)).
818 apply le_times_div_m_m.
822 rewrite < distr_times_plus.
824 apply (trans_le ? ((2*n + n)/3))
825 [apply le_plus_div.apply lt_O_S
827 change in ⊢ (? (? % ?) ?) with (3*n).
829 rewrite > lt_O_to_div_times
835 |apply (trans_le ? (2*n*B(2*n)))
837 apply (trans_le ? ? ? ? H).
838 apply leb_true_to_le.reflexivity
839 |rewrite > S_pred in ⊢ (? ? (? ? (? ? %)))
841 rewrite < assoc_times.
842 rewrite < sym_times in ⊢ (? ? (? % ?)).
843 rewrite > assoc_times in ⊢ (? ? %).
845 apply not_bertrand_to_le_B;assumption
846 |apply le_times_to_le_div
848 |apply (trans_le ? (sqrt (exp 2 8)))
849 [apply leb_true_to_le.reflexivity
850 |apply monotonic_sqrt.
851 change in ⊢ (? % ?) with (2*(exp 2 7)).
862 theorem not_bertrand_to_le2:
863 \forall n.exp 2 7 \le n \to not_bertrand n \to
864 2*n / 3 \le (sqrt(2*n)/2)*S(log 2 (2*n)).
866 rewrite < (eq_log_exp 2)
867 [apply (trans_le ? (log 2 ((exp (2*n) (sqrt(2*n)/2)))))
870 |apply not_bertrand_to_le1;assumption
879 theorem tech1: \forall a,b,c,d.O < b \to O < d \to
880 (a/b)*(c/d) \le (a*c)/(b*d).
882 apply le_times_to_le_div
883 [rewrite > (times_n_O O).
884 apply lt_times;assumption
885 |rewrite > assoc_times.
886 rewrite < assoc_times in ⊢ (? (? ? %) ?).
887 rewrite < sym_times in ⊢ (? (? ? (? % ?)) ?).
888 rewrite > assoc_times.
889 rewrite < assoc_times.
891 rewrite > sym_times;apply le_times_div_m_m;assumption
895 theorem tech: \forall n. 2*(S(log 2 (2*n))) \le sqrt (2*n) \to
896 (sqrt(2*n)/2)*S(log 2 (2*n)) \le 2*n / 4.
898 cut (4*(S(log 2 (2*n))) \le 2* sqrt(2*n))
899 [rewrite > sym_times.
900 apply le_times_to_le_div
902 |rewrite < assoc_times.
903 apply (trans_le ? (2*sqrt(2*n)*(sqrt (2*n)/2)))
904 [apply le_times_l.assumption
905 |apply (trans_le ? ((2*sqrt(2*n)*(sqrt(2*n))/2)))
906 [apply le_times_div_div_times.
908 |rewrite > assoc_times.
910 rewrite > lt_O_to_div_times.
916 |change in ⊢ (? (? % ?) ?) with (2*2).
917 rewrite > assoc_times.
923 theorem lt_div_S_div: \forall n,m. O < m \to exp m 2 \le n \to
926 apply lt_times_to_lt_div.
927 apply (lt_to_le_to_lt ? (S(n/m)*m))
928 [apply lt_div_S.assumption
929 |rewrite > sym_times in ⊢ (? ? %). simplify.
930 rewrite > sym_times in ⊢ (? ? (? ? %)).
932 apply le_times_to_le_div
940 theorem exp_plus_SSO: \forall a,b. exp (a+b) 2 = (exp a 2) + 2*a*b + (exp b 2).
943 rewrite > distr_times_plus.
944 rewrite > times_plus_l.
946 rewrite > assoc_plus.
947 rewrite > assoc_plus.
949 rewrite > times_plus_l.
951 rewrite < assoc_plus.
953 rewrite > plus_n_O in ⊢ (? ? (? (? ? %) ?) ?).
954 rewrite > assoc_times.
955 apply eq_f2;reflexivity.
958 theorem tech3: \forall n. (exp 2 8) \le n \to 2*(S(log 2 (2*n))) \le sqrt (2*n).
960 lapply (le_log 2 ? ? (le_n ?) H) as H1.
961 rewrite > exp_n_SO in ⊢ (? (? ? (? (? ? (? % ?)))) ?).
968 apply (trans_le ? (2*log 2 n))
969 [rewrite < times_SSO_n.
972 [apply leb_true_to_le.reflexivity
973 |rewrite < (eq_log_exp 2)
978 |apply (trans_le ? ? ? ? (le_exp_log 2 ? ? )).
979 apply le_times_SSO_n_exp_SSO_n.
980 apply (lt_to_le_to_lt ? ? ? ? H).
981 apply leb_true_to_le.reflexivity
983 |apply le_to_leb_true.
984 rewrite > assoc_times.
987 rewrite > assoc_times.
989 rewrite > exp_plus_SSO.
990 rewrite > distr_times_plus.
991 rewrite > distr_times_plus.
992 rewrite > assoc_plus.
993 apply (trans_le ? (4*exp (log 2 n) 2))
994 [change in ⊢ (? ? (? % ?)) with (2*2).
995 rewrite > assoc_times in ⊢ (? ? %).
996 rewrite < times_SSO_n in ⊢ (? ? %).
998 rewrite < times_SSO_n in ⊢ (? ? %).
1000 [rewrite > sym_times in ⊢ (? (? ? %) ?).
1001 rewrite < assoc_times.
1002 rewrite < assoc_times.
1003 change in ⊢ (? (? % ?) ?) with 8.
1006 (* strange things here *)
1007 rewrite < (eq_log_exp 2)
1011 |apply (trans_le ? (log 2 n))
1012 [change in ⊢ (? % ?) with 8.
1013 rewrite < (eq_log_exp 2)
1017 |rewrite > exp_n_SO in ⊢ (? % ?).
1020 [apply (lt_to_le_to_lt ? ? ? ? H).
1021 apply leb_true_to_le.reflexivity
1022 |apply (lt_to_le_to_lt ? ? ? ? H).
1023 apply leb_true_to_le.reflexivity
1029 |change in ⊢ (? (? % ?) ?) with (exp 2 2).
1030 apply (trans_le ? ? ? ? (le_exp_log 2 ? ?))
1031 [apply le_times_exp_n_SSO_exp_SSO_n
1033 |change in ⊢ (? % ?) with 8.
1034 rewrite < (eq_log_exp 2)
1039 |apply (lt_to_le_to_lt ? ? ? ? H).
1040 apply leb_true_to_le.reflexivity
1045 |apply (lt_to_le_to_lt ? ? ? ? H).
1046 apply leb_true_to_le.reflexivity
1050 theorem le_to_bertrand2:
1051 \forall n. (exp 2 8) \le n \to bertrand n.
1053 apply not_not_bertrand_to_bertrand.unfold.intro.
1054 absurd (2*n / 3 \le (sqrt(2*n)/2)*S(log 2 (2*n)))
1055 [apply not_bertrand_to_le2
1056 [apply (trans_le ? ? ? ? H).
1063 |apply lt_to_not_le.
1064 apply (le_to_lt_to_lt ? ? ? ? (lt_div_S_div ? ? ? ?))
1065 [apply tech.apply tech3.assumption
1067 |apply (trans_le ? (2*exp 2 8))
1068 [apply leb_true_to_le.reflexivity
1069 |apply le_times_r.assumption
1075 theorem bertrand_n :
1076 \forall n. O < n \to bertrand n.
1077 intros;elim (decidable_le n 256)
1078 [apply le_to_bertrand;assumption
1079 |apply le_to_bertrand2;apply lt_to_le;apply not_le_to_lt;apply H1]
1083 theorem mod_exp: eqb (mod (exp 2 8) 13) O = false.