]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/nth_prime.ma
experimental branch with no set baseuri command and no developments
[helm.git] / matita / library / nat / nth_prime.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
16
17 include "nat/primes.ma".
18 include "nat/lt_arith.ma".
19
20 (* upper bound by Bertrand's conjecture. *)
21 (* Too difficult to prove.        
22 let rec nth_prime n \def
23 match n with
24   [ O \Rightarrow (S(S O))
25   | (S p) \Rightarrow
26     let previous_prime \def S (nth_prime p) in
27     min_aux previous_prime ((S(S O))*previous_prime) primeb].
28
29 theorem example8 : nth_prime (S(S O)) = (S(S(S(S(S O))))).
30 normalize.reflexivity.
31 qed.
32
33 theorem example9 : nth_prime (S(S(S O))) = (S(S(S(S(S(S(S O))))))).
34 normalize.reflexivity.
35 qed.
36
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.
39 qed. *)
40
41 theorem smallest_factor_fact: \forall n:nat.
42 n < smallest_factor (S n!).
43 intros.
44 apply not_le_to_lt.unfold Not.
45 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.
49 assumption.
50 apply divides_smallest_factor_n.
51 unfold lt.apply le_S_S.apply le_O_n.
52 qed.
53
54 theorem ex_prime: \forall n. (S O) \le n \to \exists m.
55 n < m \land m \le S n! \land (prime m).
56 intros.
57 elim H.
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)!))).
62 split.split.
63 apply smallest_factor_fact.
64 apply le_smallest_factor_n.
65 (* Andrea: ancora hint non lo trova *)
66 apply prime_smallest_factor_n.unfold lt.
67 apply le_S.apply le_SSO_fact.
68 unfold lt.apply le_S_S.assumption.
69 qed.
70
71 let rec nth_prime n \def
72 match n with
73   [ O \Rightarrow (S(S O))
74   | (S p) \Rightarrow
75     let previous_prime \def (nth_prime p) in
76     let upper_bound \def S previous_prime! in
77     min_aux upper_bound (S previous_prime) primeb].
78     
79 (* it works
80 theorem example11 : nth_prime (S(S O)) = (S(S(S(S(S O))))).
81 normalize.reflexivity.
82 qed.
83
84 theorem example12: nth_prime (S(S(S O))) = (S(S(S(S(S(S(S O))))))).
85 normalize.reflexivity.
86 qed.
87
88 theorem example13 : nth_prime (S(S(S(S O)))) = (S(S(S(S(S(S(S(S(S(S(S O))))))))))).
89 normalize.reflexivity.
90 qed.
91
92 alias num (instance 0) = "natural number".
93 theorem example14 : nth_prime 18 = 67.
94 normalize.reflexivity.
95 qed.
96 *)
97
98 theorem prime_nth_prime : \forall n:nat.prime (nth_prime n).
99 intro.
100 apply (nat_case n).simplify.
101 apply (primeb_to_Prop (S(S O))).
102 intro.
103 change with
104 (let previous_prime \def (nth_prime m) in
105 let upper_bound \def S previous_prime! in
106 prime (min_aux upper_bound (S previous_prime) primeb)).
107 apply primeb_true_to_prime.
108 apply f_min_aux_true.
109 apply (ex_intro nat ? (smallest_factor (S (nth_prime m)!))).
110 split.split.
111 apply smallest_factor_fact.
112 apply transitive_le;
113  [2: apply le_smallest_factor_n
114  | skip
115  | apply (le_plus_n_r (S (nth_prime m)) (S (fact (nth_prime m))))
116  ].
117 apply prime_to_primeb_true.
118 apply prime_smallest_factor_n.unfold lt.
119 apply le_S_S.apply le_SO_fact.
120 qed.
121
122 (* properties of nth_prime *)
123 theorem increasing_nth_prime: increasing nth_prime.
124 unfold increasing.
125 intros.
126 change with
127 (let previous_prime \def (nth_prime n) in
128 let upper_bound \def S previous_prime! in
129 (S previous_prime) \le min_aux upper_bound (S previous_prime) primeb).
130 intros.
131 apply le_min_aux.
132 qed.
133
134 variant lt_nth_prime_n_nth_prime_Sn :\forall n:nat. 
135 (nth_prime n) < (nth_prime (S n)) \def increasing_nth_prime.
136
137 theorem injective_nth_prime: injective nat nat nth_prime.
138 apply increasing_to_injective.
139 apply increasing_nth_prime.
140 qed.
141
142 theorem lt_SO_nth_prime_n : \forall n:nat. (S O) \lt nth_prime n.
143 intros. elim n.unfold lt.apply le_n.
144 apply (trans_lt ? (nth_prime n1)).
145 assumption.apply lt_nth_prime_n_nth_prime_Sn.
146 qed.
147
148 theorem lt_O_nth_prime_n : \forall n:nat. O \lt nth_prime n.
149 intros.apply (trans_lt O (S O)).
150 unfold lt. apply le_n.apply lt_SO_nth_prime_n.
151 qed.
152
153 theorem lt_n_nth_prime_n : \forall n:nat. n \lt nth_prime n.
154 intro.
155 elim n
156   [apply lt_O_nth_prime_n
157   |apply (lt_to_le_to_lt ? (S (nth_prime n1)))
158     [unfold.apply le_S_S.assumption
159     |apply lt_nth_prime_n_nth_prime_Sn
160     ]
161   ]
162 qed.
163
164 theorem ex_m_le_n_nth_prime_m: 
165 \forall n: nat. nth_prime O \le n \to 
166 \exists m. nth_prime m \le n \land n < nth_prime (S m).
167 intros.
168 apply increasing_to_le2.
169 exact lt_nth_prime_n_nth_prime_Sn.assumption.
170 qed.
171
172 theorem lt_nth_prime_to_not_prime: \forall n,m. nth_prime n < m \to m < nth_prime (S n) 
173 \to \lnot (prime m).
174 intros.
175 apply primeb_false_to_not_prime.
176 letin previous_prime \def (nth_prime n).
177 letin upper_bound \def (S previous_prime!).
178 apply (lt_min_aux_to_false primeb (S previous_prime) upper_bound m).
179 assumption.
180 unfold lt.
181 apply (transitive_le (S m) (nth_prime (S n)) (min_aux (S (fact (nth_prime n))) (S (nth_prime n)) primeb) ? ?);
182   [apply (H1).
183   |apply (le_n (min_aux (S (fact (nth_prime n))) (S (nth_prime n)) primeb)).
184   ]
185 qed.
186
187 (* nth_prime enumerates all primes *)
188 theorem prime_to_nth_prime : \forall p:nat. prime p \to
189 \exists i. nth_prime i = p.
190 intros.
191 cut (\exists m. nth_prime m \le p \land p < nth_prime (S m)).
192 elim Hcut.elim H1.
193 cut (nth_prime a < p \lor nth_prime a = p).
194 elim Hcut1.
195 absurd (prime p).
196 assumption.
197 apply (lt_nth_prime_to_not_prime a).assumption.assumption.
198 apply (ex_intro nat ? a).assumption.
199 apply le_to_or_lt_eq.assumption.
200 apply ex_m_le_n_nth_prime_m.
201 simplify.unfold prime in H.elim H.assumption.
202 qed.
203