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 set "baseuri" "cic:/matita/nat/nth_prime".
17 include "nat/primes.ma".
18 include "nat/lt_arith.ma".
20 (* upper bound by Bertrand's conjecture. *)
21 (* Too difficult to prove.
22 let rec nth_prime n \def
24 [ O \Rightarrow (S(S O))
26 let previous_prime \def S (nth_prime p) in
27 min_aux previous_prime ((S(S O))*previous_prime) primeb].
29 theorem example8 : nth_prime (S(S O)) = (S(S(S(S(S O))))).
30 normalize.reflexivity.
33 theorem example9 : nth_prime (S(S(S O))) = (S(S(S(S(S(S(S O))))))).
34 normalize.reflexivity.
37 theorem example10 : nth_prime (S(S(S(S O)))) = (S(S(S(S(S(S(S(S(S(S(S O))))))))))).
38 normalize.reflexivity.
41 theorem smallest_factor_fact: \forall n:nat.
42 n < smallest_factor (S n!).
45 change with (smallest_factor (S n!) \le n \to False).intro.
46 apply (not_divides_S_fact n (smallest_factor(S n!))).
47 apply lt_SO_smallest_factor.
48 unfold lt.apply le_S_S.apply le_SO_fact.
50 apply divides_smallest_factor_n.
51 unfold lt.apply le_S_S.apply le_O_n.
54 theorem ex_prime: \forall n. (S O) \le n \to \exists m.
55 n < m \land m \le S n! \land (prime m).
58 apply (ex_intro nat ? (S(S O))).
59 split.split.apply (le_n (S(S O))).
60 apply (le_n (S(S O))).apply (primeb_to_Prop (S(S O))).
61 apply (ex_intro nat ? (smallest_factor (S (S n1)!))).
63 apply smallest_factor_fact.
64 apply le_smallest_factor_n.
65 (* Andrea: ancora hint non lo trova *)
66 apply prime_smallest_factor_n.
67 change with ((S(S O)) \le S (S n1)!).
68 apply le_S.apply le_SSO_fact.
69 unfold lt.apply le_S_S.assumption.
72 let rec nth_prime n \def
74 [ O \Rightarrow (S(S O))
76 let previous_prime \def (nth_prime p) in
77 let upper_bound \def S previous_prime! in
78 min_aux (upper_bound - (S previous_prime)) upper_bound primeb].
80 (* it works, but nth_prime 4 takes already a few minutes -
81 it must compute factorial of 7 ...
83 theorem example11 : nth_prime (S(S O)) = (S(S(S(S(S O))))).
84 normalize.reflexivity.
87 theorem example12: nth_prime (S(S(S O))) = (S(S(S(S(S(S(S O))))))).
88 normalize.reflexivity.
91 theorem example13 : nth_prime (S(S(S(S O)))) = (S(S(S(S(S(S(S(S(S(S(S O))))))))))).
92 normalize.reflexivity.
95 theorem prime_nth_prime : \forall n:nat.prime (nth_prime n).
98 change with (prime (S(S O))).
99 apply (primeb_to_Prop (S(S O))).
102 (let previous_prime \def (nth_prime m) in
103 let upper_bound \def S previous_prime! in
104 prime (min_aux (upper_bound - (S previous_prime)) upper_bound primeb)).
105 apply primeb_true_to_prime.
106 apply f_min_aux_true.
107 apply (ex_intro nat ? (smallest_factor (S (nth_prime m)!))).
109 cut (S (nth_prime m)!-(S (nth_prime m)! - (S (nth_prime m))) = (S (nth_prime m))).
110 rewrite > Hcut.exact (smallest_factor_fact (nth_prime m)).
111 (* maybe we could factorize this proof *)
113 apply plus_minus_m_m.
116 apply le_smallest_factor_n.
117 apply prime_to_primeb_true.
118 apply prime_smallest_factor_n.
119 change with ((S(S O)) \le S (nth_prime m)!).
120 apply le_S_S.apply le_SO_fact.
123 (* properties of nth_prime *)
124 theorem increasing_nth_prime: increasing nth_prime.
125 change with (\forall n:nat. (nth_prime n) < (nth_prime (S n))).
128 (let previous_prime \def (nth_prime n) in
129 let upper_bound \def S previous_prime! in
130 (S previous_prime) \le min_aux (upper_bound - (S previous_prime)) upper_bound primeb).
132 cut (upper_bound - (upper_bound -(S previous_prime)) = (S previous_prime)).
133 rewrite < Hcut in \vdash (? % ?).
136 apply plus_minus_m_m.
141 variant lt_nth_prime_n_nth_prime_Sn :\forall n:nat.
142 (nth_prime n) < (nth_prime (S n)) \def increasing_nth_prime.
144 theorem injective_nth_prime: injective nat nat nth_prime.
145 apply increasing_to_injective.
146 apply increasing_nth_prime.
149 theorem lt_SO_nth_prime_n : \forall n:nat. (S O) \lt nth_prime n.
150 intros. elim n.unfold lt.apply le_n.
151 apply (trans_lt ? (nth_prime n1)).
152 assumption.apply lt_nth_prime_n_nth_prime_Sn.
155 theorem lt_O_nth_prime_n : \forall n:nat. O \lt nth_prime n.
156 intros.apply (trans_lt O (S O)).
157 unfold lt. apply le_n.apply lt_SO_nth_prime_n.
160 theorem ex_m_le_n_nth_prime_m:
161 \forall n: nat. nth_prime O \le n \to
162 \exists m. nth_prime m \le n \land n < nth_prime (S m).
164 apply increasing_to_le2.
165 exact lt_nth_prime_n_nth_prime_Sn.assumption.
168 theorem lt_nth_prime_to_not_prime: \forall n,m. nth_prime n < m \to m < nth_prime (S n)
171 apply primeb_false_to_not_prime.
172 letin previous_prime \def (nth_prime n).
173 letin upper_bound \def (S previous_prime!).
174 apply (lt_min_aux_to_false primeb upper_bound (upper_bound - (S previous_prime)) m).
175 cut (S (nth_prime n)!-(S (nth_prime n)! - (S (nth_prime n))) = (S (nth_prime n))).
176 rewrite > Hcut.assumption.
178 apply plus_minus_m_m.
184 (* nth_prime enumerates all primes *)
185 theorem prime_to_nth_prime : \forall p:nat. prime p \to
186 \exists i. nth_prime i = p.
188 cut (\exists m. nth_prime m \le p \land p < nth_prime (S m)).
190 cut (nth_prime a < p \lor nth_prime a = p).
194 apply (lt_nth_prime_to_not_prime a).assumption.assumption.
195 apply (ex_intro nat ? a).assumption.
196 apply le_to_or_lt_eq.assumption.
197 apply ex_m_le_n_nth_prime_m.
198 simplify.unfold prime in H.elim H.assumption.