author | wenzelm |
Wed, 04 Oct 2017 12:00:53 +0200 | |
changeset 66787 | 64b47495676d |
parent 66453 | cc19f7ca2ed6 |
child 70749 | 5d06b7bb9d22 |
permissions | -rw-r--r-- |
27468 | 1 |
(* Title : NSPrimes.thy |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 2002 University of Edinburgh |
|
4 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2004 |
|
5 |
*) |
|
6 |
||
64601 | 7 |
section \<open>The Nonstandard Primes as an Extension of the Prime Numbers\<close> |
27468 | 8 |
|
9 |
theory NSPrimes |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
65417
diff
changeset
|
10 |
imports "HOL-Computational_Algebra.Primes" "HOL-Nonstandard_Analysis.Hyperreal" |
27468 | 11 |
begin |
12 |
||
64601 | 13 |
text \<open>These can be used to derive an alternative proof of the infinitude of |
61975 | 14 |
primes by considering a property of nonstandard sets.\<close> |
27468 | 15 |
|
64601 | 16 |
definition starprime :: "hypnat set" |
17 |
where [transfer_unfold]: "starprime = *s* {p. prime p}" |
|
27468 | 18 |
|
64601 | 19 |
definition choicefun :: "'a set \<Rightarrow> 'a" |
20 |
where "choicefun E = (SOME x. \<exists>X \<in> Pow E - {{}}. x \<in> X)" |
|
27468 | 21 |
|
64601 | 22 |
primrec injf_max :: "nat \<Rightarrow> 'a::order set \<Rightarrow> 'a" |
48488 | 23 |
where |
27468 | 24 |
injf_max_zero: "injf_max 0 E = choicefun E" |
64601 | 25 |
| injf_max_Suc: "injf_max (Suc n) E = choicefun ({e. e \<in> E \<and> injf_max n E < e})" |
27468 | 26 |
|
64601 | 27 |
lemma dvd_by_all2: "\<exists>N>0. \<forall>m. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N" |
28 |
for M :: nat |
|
29 |
apply (induct M) |
|
30 |
apply auto |
|
31 |
apply (rule_tac x = "N * Suc M" in exI) |
|
32 |
apply auto |
|
33 |
apply (metis dvdI dvd_add_times_triv_left_iff dvd_add_triv_right_iff dvd_refl dvd_trans le_Suc_eq mult_Suc_right) |
|
34 |
done |
|
27468 | 35 |
|
64601 | 36 |
lemma dvd_by_all: "\<forall>M::nat. \<exists>N>0. \<forall>m. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N" |
62349
7c23469b5118
cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents:
61975
diff
changeset
|
37 |
using dvd_by_all2 by blast |
27468 | 38 |
|
64601 | 39 |
lemma hypnat_of_nat_le_zero_iff [simp]: "hypnat_of_nat n \<le> 0 \<longleftrightarrow> n = 0" |
40 |
by transfer simp |
|
27468 | 41 |
|
64601 | 42 |
text \<open>Goldblatt: Exercise 5.11(2) -- p. 57.\<close> |
43 |
lemma hdvd_by_all: "\<forall>M. \<exists>N. 0 < N \<and> (\<forall>m::hypnat. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N)" |
|
44 |
by transfer (rule dvd_by_all) |
|
27468 | 45 |
|
45605 | 46 |
lemmas hdvd_by_all2 = hdvd_by_all [THEN spec] |
27468 | 47 |
|
64601 | 48 |
text \<open>Goldblatt: Exercise 5.11(2) -- p. 57.\<close> |
27468 | 49 |
lemma hypnat_dvd_all_hypnat_of_nat: |
64601 | 50 |
"\<exists>N::hypnat. 0 < N \<and> (\<forall>n \<in> - {0::nat}. hypnat_of_nat n dvd N)" |
51 |
apply (cut_tac hdvd_by_all) |
|
52 |
apply (drule_tac x = whn in spec) |
|
53 |
apply auto |
|
54 |
apply (rule exI) |
|
55 |
apply auto |
|
56 |
apply (drule_tac x = "hypnat_of_nat n" in spec) |
|
57 |
apply (auto simp add: linorder_not_less) |
|
58 |
done |
|
27468 | 59 |
|
60 |
||
64601 | 61 |
text \<open>The nonstandard extension of the set prime numbers consists of precisely |
62 |
those hypernaturals exceeding 1 that have no nontrivial factors.\<close> |
|
27468 | 63 |
|
64601 | 64 |
text \<open>Goldblatt: Exercise 5.11(3a) -- p 57.\<close> |
65 |
lemma starprime: "starprime = {p. 1 < p \<and> (\<forall>m. m dvd p \<longrightarrow> m = 1 \<or> m = p)}" |
|
66 |
by transfer (auto simp add: prime_nat_iff) |
|
27468 | 67 |
|
64601 | 68 |
text \<open>Goldblatt Exercise 5.11(3b) -- p 57.\<close> |
69 |
lemma hyperprime_factor_exists: "\<And>n. 1 < n \<Longrightarrow> \<exists>k \<in> starprime. k dvd n" |
|
70 |
by transfer (simp add: prime_factor_nat) |
|
27468 | 71 |
|
64601 | 72 |
text \<open>Goldblatt Exercise 3.10(1) -- p. 29.\<close> |
73 |
lemma NatStar_hypnat_of_nat: "finite A \<Longrightarrow> *s* A = hypnat_of_nat ` A" |
|
74 |
by (rule starset_finite) |
|
27468 | 75 |
|
76 |
||
64601 | 77 |
subsection \<open>Another characterization of infinite set of natural numbers\<close> |
27468 | 78 |
|
64601 | 79 |
lemma finite_nat_set_bounded: "finite N \<Longrightarrow> \<exists>n::nat. \<forall>i \<in> N. i < n" |
80 |
apply (erule_tac F = N in finite_induct) |
|
81 |
apply auto |
|
82 |
apply (rule_tac x = "Suc n + x" in exI) |
|
83 |
apply auto |
|
84 |
done |
|
27468 | 85 |
|
64601 | 86 |
lemma finite_nat_set_bounded_iff: "finite N \<longleftrightarrow> (\<exists>n::nat. \<forall>i \<in> N. i < n)" |
87 |
by (blast intro: finite_nat_set_bounded bounded_nat_set_is_finite) |
|
27468 | 88 |
|
64601 | 89 |
lemma not_finite_nat_set_iff: "\<not> finite N \<longleftrightarrow> (\<forall>n::nat. \<exists>i \<in> N. n \<le> i)" |
90 |
by (auto simp add: finite_nat_set_bounded_iff not_less) |
|
27468 | 91 |
|
64601 | 92 |
lemma bounded_nat_set_is_finite2: "\<forall>i::nat \<in> N. i \<le> n \<Longrightarrow> finite N" |
93 |
apply (rule finite_subset) |
|
94 |
apply (rule_tac [2] finite_atMost) |
|
95 |
apply auto |
|
96 |
done |
|
27468 | 97 |
|
64601 | 98 |
lemma finite_nat_set_bounded2: "finite N \<Longrightarrow> \<exists>n::nat. \<forall>i \<in> N. i \<le> n" |
99 |
apply (erule_tac F = N in finite_induct) |
|
100 |
apply auto |
|
101 |
apply (rule_tac x = "n + x" in exI) |
|
102 |
apply auto |
|
103 |
done |
|
27468 | 104 |
|
64601 | 105 |
lemma finite_nat_set_bounded_iff2: "finite N \<longleftrightarrow> (\<exists>n::nat. \<forall>i \<in> N. i \<le> n)" |
106 |
by (blast intro: finite_nat_set_bounded2 bounded_nat_set_is_finite2) |
|
27468 | 107 |
|
64601 | 108 |
lemma not_finite_nat_set_iff2: "\<not> finite N \<longleftrightarrow> (\<forall>n::nat. \<exists>i \<in> N. n < i)" |
109 |
by (auto simp add: finite_nat_set_bounded_iff2 not_le) |
|
27468 | 110 |
|
111 |
||
64601 | 112 |
subsection \<open>An injective function cannot define an embedded natural number\<close> |
27468 | 113 |
|
64601 | 114 |
lemma lemma_infinite_set_singleton: |
115 |
"\<forall>m n. m \<noteq> n \<longrightarrow> f n \<noteq> f m \<Longrightarrow> {n. f n = N} = {} \<or> (\<exists>m. {n. f n = N} = {m})" |
|
116 |
apply auto |
|
117 |
apply (drule_tac x = x in spec, auto) |
|
118 |
apply (subgoal_tac "\<forall>n. f n = f x \<longleftrightarrow> x = n") |
|
119 |
apply auto |
|
120 |
done |
|
27468 | 121 |
|
122 |
lemma inj_fun_not_hypnat_in_SHNat: |
|
64601 | 123 |
fixes f :: "nat \<Rightarrow> nat" |
124 |
assumes inj_f: "inj f" |
|
27468 | 125 |
shows "starfun f whn \<notin> Nats" |
126 |
proof |
|
127 |
from inj_f have inj_f': "inj (starfun f)" |
|
128 |
by (transfer inj_on_def Ball_def UNIV_def) |
|
129 |
assume "starfun f whn \<in> Nats" |
|
130 |
then obtain N where N: "starfun f whn = hypnat_of_nat N" |
|
64601 | 131 |
by (auto simp: Nats_def) |
132 |
then have "\<exists>n. starfun f n = hypnat_of_nat N" .. |
|
133 |
then have "\<exists>n. f n = N" by transfer |
|
134 |
then obtain n where "f n = N" .. |
|
135 |
then have "starfun f (hypnat_of_nat n) = hypnat_of_nat N" |
|
27468 | 136 |
by transfer |
137 |
with N have "starfun f whn = starfun f (hypnat_of_nat n)" |
|
138 |
by simp |
|
139 |
with inj_f' have "whn = hypnat_of_nat n" |
|
140 |
by (rule injD) |
|
64601 | 141 |
then show False |
27468 | 142 |
by (simp add: whn_neq_hypnat_of_nat) |
143 |
qed |
|
144 |
||
64601 | 145 |
lemma range_subset_mem_starsetNat: "range f \<subseteq> A \<Longrightarrow> starfun f whn \<in> *s* A" |
146 |
apply (rule_tac x="whn" in spec) |
|
147 |
apply transfer |
|
148 |
apply auto |
|
149 |
done |
|
150 |
||
151 |
text \<open> |
|
152 |
Gleason Proposition 11-5.5. pg 149, pg 155 (ex. 3) and pg. 360. |
|
27468 | 153 |
|
64601 | 154 |
Let \<open>E\<close> be a nonvoid ordered set with no maximal elements (note: effectively an |
155 |
infinite set if we take \<open>E = N\<close> (Nats)). Then there exists an order-preserving |
|
156 |
injection from \<open>N\<close> to \<open>E\<close>. Of course, (as some doofus will undoubtedly point out! |
|
157 |
:-)) can use notion of least element in proof (i.e. no need for choice) if |
|
158 |
dealing with nats as we have well-ordering property. |
|
159 |
\<close> |
|
27468 | 160 |
|
64601 | 161 |
lemma lemmaPow3: "E \<noteq> {} \<Longrightarrow> \<exists>x. \<exists>X \<in> Pow E - {{}}. x \<in> X" |
162 |
by auto |
|
27468 | 163 |
|
64601 | 164 |
lemma choicefun_mem_set [simp]: "E \<noteq> {} \<Longrightarrow> choicefun E \<in> E" |
165 |
apply (unfold choicefun_def) |
|
166 |
apply (rule lemmaPow3 [THEN someI2_ex], auto) |
|
167 |
done |
|
27468 | 168 |
|
64601 | 169 |
lemma injf_max_mem_set: "E \<noteq>{} \<Longrightarrow> \<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> injf_max n E \<in> E" |
170 |
apply (induct n) |
|
171 |
apply force |
|
172 |
apply (simp add: choicefun_def) |
|
173 |
apply (rule lemmaPow3 [THEN someI2_ex], auto) |
|
174 |
done |
|
27468 | 175 |
|
64601 | 176 |
lemma injf_max_order_preserving: "\<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> injf_max n E < injf_max (Suc n) E" |
177 |
apply (simp add: choicefun_def) |
|
178 |
apply (rule lemmaPow3 [THEN someI2_ex]) |
|
179 |
apply auto |
|
180 |
done |
|
27468 | 181 |
|
64601 | 182 |
lemma injf_max_order_preserving2: "\<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> \<forall>n m. m < n \<longrightarrow> injf_max m E < injf_max n E" |
183 |
apply (rule allI) |
|
184 |
apply (induct_tac n) |
|
185 |
apply auto |
|
186 |
apply (simp add: choicefun_def) |
|
187 |
apply (rule lemmaPow3 [THEN someI2_ex]) |
|
188 |
apply (auto simp add: less_Suc_eq) |
|
189 |
apply (drule_tac x = m in spec) |
|
190 |
apply (drule subsetD) |
|
191 |
apply auto |
|
192 |
apply (drule_tac x = "injf_max m E" in order_less_trans) |
|
193 |
apply auto |
|
194 |
done |
|
27468 | 195 |
|
64601 | 196 |
lemma inj_injf_max: "\<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> inj (\<lambda>n. injf_max n E)" |
197 |
apply (rule inj_onI) |
|
198 |
apply (rule ccontr) |
|
199 |
apply auto |
|
200 |
apply (drule injf_max_order_preserving2) |
|
201 |
apply (metis linorder_antisym_conv3 order_less_le) |
|
202 |
done |
|
27468 | 203 |
|
204 |
lemma infinite_set_has_order_preserving_inj: |
|
64601 | 205 |
"E \<noteq> {} \<Longrightarrow> \<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> \<exists>f. range f \<subseteq> E \<and> inj f \<and> (\<forall>m. f m < f (Suc m))" |
206 |
for E :: "'a::order set" and f :: "nat \<Rightarrow> 'a" |
|
207 |
apply (rule_tac x = "\<lambda>n. injf_max n E" in exI) |
|
208 |
apply safe |
|
209 |
apply (rule injf_max_mem_set) |
|
210 |
apply (rule_tac [3] inj_injf_max) |
|
211 |
apply (rule_tac [4] injf_max_order_preserving) |
|
212 |
apply auto |
|
213 |
done |
|
27468 | 214 |
|
215 |
||
64601 | 216 |
text \<open>Only need the existence of an injective function from \<open>N\<close> to \<open>A\<close> for proof.\<close> |
27468 | 217 |
|
64601 | 218 |
lemma hypnat_infinite_has_nonstandard: "\<not> finite A \<Longrightarrow> hypnat_of_nat ` A < ( *s* A)" |
219 |
apply auto |
|
220 |
apply (subgoal_tac "A \<noteq> {}") |
|
221 |
prefer 2 apply force |
|
222 |
apply (drule infinite_set_has_order_preserving_inj) |
|
223 |
apply (erule not_finite_nat_set_iff2 [THEN iffD1]) |
|
224 |
apply auto |
|
225 |
apply (drule inj_fun_not_hypnat_in_SHNat) |
|
226 |
apply (drule range_subset_mem_starsetNat) |
|
227 |
apply (auto simp add: SHNat_eq) |
|
228 |
done |
|
27468 | 229 |
|
64601 | 230 |
lemma starsetNat_eq_hypnat_of_nat_image_finite: "*s* A = hypnat_of_nat ` A \<Longrightarrow> finite A" |
231 |
by (metis hypnat_infinite_has_nonstandard less_irrefl) |
|
232 |
||
233 |
lemma finite_starsetNat_iff: "*s* A = hypnat_of_nat ` A \<longleftrightarrow> finite A" |
|
234 |
by (blast intro!: starsetNat_eq_hypnat_of_nat_image_finite NatStar_hypnat_of_nat) |
|
27468 | 235 |
|
64601 | 236 |
lemma hypnat_infinite_has_nonstandard_iff: "\<not> finite A \<longleftrightarrow> hypnat_of_nat ` A < *s* A" |
237 |
apply (rule iffI) |
|
238 |
apply (blast intro!: hypnat_infinite_has_nonstandard) |
|
239 |
apply (auto simp add: finite_starsetNat_iff [symmetric]) |
|
240 |
done |
|
27468 | 241 |
|
242 |
||
64601 | 243 |
subsection \<open>Existence of Infinitely Many Primes: a Nonstandard Proof\<close> |
27468 | 244 |
|
64601 | 245 |
lemma lemma_not_dvd_hypnat_one [simp]: "\<not> (\<forall>n \<in> - {0}. hypnat_of_nat n dvd 1)" |
246 |
apply auto |
|
247 |
apply (rule_tac x = 2 in bexI) |
|
248 |
apply transfer |
|
249 |
apply auto |
|
250 |
done |
|
27468 | 251 |
|
64601 | 252 |
lemma lemma_not_dvd_hypnat_one2 [simp]: "\<exists>n \<in> - {0}. \<not> hypnat_of_nat n dvd 1" |
253 |
using lemma_not_dvd_hypnat_one by (auto simp del: lemma_not_dvd_hypnat_one) |
|
254 |
||
255 |
lemma hypnat_add_one_gt_one: "\<And>N::hypnat. 0 < N \<Longrightarrow> 1 < N + 1" |
|
256 |
by transfer simp |
|
27468 | 257 |
|
55165
f4791db20067
Removed a dependence upon Old_Number_Theory. Simplified a few proofs.
paulson <lp15@cam.ac.uk>
parents:
48488
diff
changeset
|
258 |
lemma hypnat_of_nat_zero_not_prime [simp]: "hypnat_of_nat 0 \<notin> starprime" |
64601 | 259 |
by transfer simp |
27468 | 260 |
|
64601 | 261 |
lemma hypnat_zero_not_prime [simp]: "0 \<notin> starprime" |
262 |
using hypnat_of_nat_zero_not_prime by simp |
|
27468 | 263 |
|
55165
f4791db20067
Removed a dependence upon Old_Number_Theory. Simplified a few proofs.
paulson <lp15@cam.ac.uk>
parents:
48488
diff
changeset
|
264 |
lemma hypnat_of_nat_one_not_prime [simp]: "hypnat_of_nat 1 \<notin> starprime" |
64601 | 265 |
by transfer simp |
27468 | 266 |
|
55165
f4791db20067
Removed a dependence upon Old_Number_Theory. Simplified a few proofs.
paulson <lp15@cam.ac.uk>
parents:
48488
diff
changeset
|
267 |
lemma hypnat_one_not_prime [simp]: "1 \<notin> starprime" |
64601 | 268 |
using hypnat_of_nat_one_not_prime by simp |
27468 | 269 |
|
64601 | 270 |
lemma hdvd_diff: "\<And>k m n :: hypnat. k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m - n)" |
271 |
by transfer (rule dvd_diff_nat) |
|
27468 | 272 |
|
64601 | 273 |
lemma hdvd_one_eq_one: "\<And>x::hypnat. is_unit x \<Longrightarrow> x = 1" |
62349
7c23469b5118
cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents:
61975
diff
changeset
|
274 |
by transfer simp |
27468 | 275 |
|
64601 | 276 |
text \<open>Already proved as \<open>primes_infinite\<close>, but now using non-standard naturals.\<close> |
277 |
theorem not_finite_prime: "\<not> finite {p::nat. prime p}" |
|
278 |
apply (rule hypnat_infinite_has_nonstandard_iff [THEN iffD2]) |
|
279 |
using hypnat_dvd_all_hypnat_of_nat |
|
280 |
apply clarify |
|
281 |
apply (drule hypnat_add_one_gt_one) |
|
282 |
apply (drule hyperprime_factor_exists) |
|
283 |
apply clarify |
|
284 |
apply (subgoal_tac "k \<notin> hypnat_of_nat ` {p. prime p}") |
|
285 |
apply (force simp: starprime_def) |
|
286 |
apply (metis Compl_iff add.commute dvd_add_left_iff empty_iff hdvd_one_eq_one hypnat_one_not_prime |
|
287 |
imageE insert_iff mem_Collect_eq not_prime_0) |
|
288 |
done |
|
27468 | 289 |
|
290 |
end |