author | wenzelm |
Mon, 11 Sep 2023 19:30:48 +0200 | |
changeset 78659 | b5f3d1051b13 |
parent 75867 | d7595b12b9ea |
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 |
|
75867 | 29 |
proof (induct M) |
30 |
case 0 |
|
31 |
then show ?case |
|
32 |
by auto |
|
33 |
next |
|
34 |
case (Suc M) |
|
35 |
then obtain N where "N>0" and "\<And>m. 0 < m \<and> m \<le> M \<Longrightarrow> m dvd N" |
|
36 |
by metis |
|
37 |
then show ?case |
|
38 |
by (metis nat_0_less_mult_iff zero_less_Suc dvd_mult dvd_mult2 dvd_refl le_Suc_eq) |
|
39 |
qed |
|
27468 | 40 |
|
64601 | 41 |
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
|
42 |
using dvd_by_all2 by blast |
27468 | 43 |
|
64601 | 44 |
lemma hypnat_of_nat_le_zero_iff [simp]: "hypnat_of_nat n \<le> 0 \<longleftrightarrow> n = 0" |
45 |
by transfer simp |
|
27468 | 46 |
|
64601 | 47 |
text \<open>Goldblatt: Exercise 5.11(2) -- p. 57.\<close> |
75867 | 48 |
lemma hdvd_by_all [rule_format]: "\<forall>M. \<exists>N. 0 < N \<and> (\<forall>m::hypnat. 0 < m \<and> m \<le> M \<longrightarrow> m dvd N)" |
64601 | 49 |
by transfer (rule dvd_by_all) |
27468 | 50 |
|
64601 | 51 |
text \<open>Goldblatt: Exercise 5.11(2) -- p. 57.\<close> |
27468 | 52 |
lemma hypnat_dvd_all_hypnat_of_nat: |
64601 | 53 |
"\<exists>N::hypnat. 0 < N \<and> (\<forall>n \<in> - {0::nat}. hypnat_of_nat n dvd N)" |
75867 | 54 |
by (metis Compl_iff gr0I hdvd_by_all hypnat_of_nat_le_whn singletonI star_of_0_less) |
27468 | 55 |
|
56 |
||
64601 | 57 |
text \<open>The nonstandard extension of the set prime numbers consists of precisely |
58 |
those hypernaturals exceeding 1 that have no nontrivial factors.\<close> |
|
27468 | 59 |
|
64601 | 60 |
text \<open>Goldblatt: Exercise 5.11(3a) -- p 57.\<close> |
61 |
lemma starprime: "starprime = {p. 1 < p \<and> (\<forall>m. m dvd p \<longrightarrow> m = 1 \<or> m = p)}" |
|
62 |
by transfer (auto simp add: prime_nat_iff) |
|
27468 | 63 |
|
64601 | 64 |
text \<open>Goldblatt Exercise 5.11(3b) -- p 57.\<close> |
65 |
lemma hyperprime_factor_exists: "\<And>n. 1 < n \<Longrightarrow> \<exists>k \<in> starprime. k dvd n" |
|
66 |
by transfer (simp add: prime_factor_nat) |
|
27468 | 67 |
|
64601 | 68 |
text \<open>Goldblatt Exercise 3.10(1) -- p. 29.\<close> |
69 |
lemma NatStar_hypnat_of_nat: "finite A \<Longrightarrow> *s* A = hypnat_of_nat ` A" |
|
70 |
by (rule starset_finite) |
|
27468 | 71 |
|
72 |
||
73 |
||
64601 | 74 |
subsection \<open>An injective function cannot define an embedded natural number\<close> |
27468 | 75 |
|
64601 | 76 |
lemma lemma_infinite_set_singleton: |
77 |
"\<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})" |
|
75867 | 78 |
by (metis (mono_tags) is_singletonI' is_singleton_the_elem mem_Collect_eq) |
27468 | 79 |
|
80 |
lemma inj_fun_not_hypnat_in_SHNat: |
|
64601 | 81 |
fixes f :: "nat \<Rightarrow> nat" |
82 |
assumes inj_f: "inj f" |
|
27468 | 83 |
shows "starfun f whn \<notin> Nats" |
84 |
proof |
|
85 |
from inj_f have inj_f': "inj (starfun f)" |
|
86 |
by (transfer inj_on_def Ball_def UNIV_def) |
|
87 |
assume "starfun f whn \<in> Nats" |
|
88 |
then obtain N where N: "starfun f whn = hypnat_of_nat N" |
|
64601 | 89 |
by (auto simp: Nats_def) |
90 |
then have "\<exists>n. starfun f n = hypnat_of_nat N" .. |
|
91 |
then have "\<exists>n. f n = N" by transfer |
|
92 |
then obtain n where "f n = N" .. |
|
93 |
then have "starfun f (hypnat_of_nat n) = hypnat_of_nat N" |
|
27468 | 94 |
by transfer |
95 |
with N have "starfun f whn = starfun f (hypnat_of_nat n)" |
|
96 |
by simp |
|
97 |
with inj_f' have "whn = hypnat_of_nat n" |
|
98 |
by (rule injD) |
|
64601 | 99 |
then show False |
27468 | 100 |
by (simp add: whn_neq_hypnat_of_nat) |
101 |
qed |
|
102 |
||
64601 | 103 |
lemma range_subset_mem_starsetNat: "range f \<subseteq> A \<Longrightarrow> starfun f whn \<in> *s* A" |
75867 | 104 |
by (metis STAR_subset_closed UNIV_I image_eqI starset_UNIV starset_image) |
64601 | 105 |
|
106 |
text \<open> |
|
107 |
Gleason Proposition 11-5.5. pg 149, pg 155 (ex. 3) and pg. 360. |
|
27468 | 108 |
|
64601 | 109 |
Let \<open>E\<close> be a nonvoid ordered set with no maximal elements (note: effectively an |
110 |
infinite set if we take \<open>E = N\<close> (Nats)). Then there exists an order-preserving |
|
111 |
injection from \<open>N\<close> to \<open>E\<close>. Of course, (as some doofus will undoubtedly point out! |
|
112 |
:-)) can use notion of least element in proof (i.e. no need for choice) if |
|
113 |
dealing with nats as we have well-ordering property. |
|
114 |
\<close> |
|
27468 | 115 |
|
64601 | 116 |
lemma lemmaPow3: "E \<noteq> {} \<Longrightarrow> \<exists>x. \<exists>X \<in> Pow E - {{}}. x \<in> X" |
117 |
by auto |
|
27468 | 118 |
|
64601 | 119 |
lemma choicefun_mem_set [simp]: "E \<noteq> {} \<Longrightarrow> choicefun E \<in> E" |
75867 | 120 |
unfolding choicefun_def |
121 |
by (force intro: lemmaPow3 [THEN someI2_ex]) |
|
27468 | 122 |
|
64601 | 123 |
lemma injf_max_mem_set: "E \<noteq>{} \<Longrightarrow> \<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> injf_max n E \<in> E" |
75867 | 124 |
proof (induct n) |
125 |
case 0 |
|
126 |
then show ?case by force |
|
127 |
next |
|
128 |
case (Suc n) |
|
129 |
then show ?case |
|
130 |
apply (simp add: choicefun_def) |
|
131 |
apply (rule lemmaPow3 [THEN someI2_ex], auto) |
|
132 |
done |
|
133 |
qed |
|
27468 | 134 |
|
64601 | 135 |
lemma injf_max_order_preserving: "\<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> injf_max n E < injf_max (Suc n) E" |
75867 | 136 |
by (metis (no_types, lifting) choicefun_mem_set empty_iff injf_max.simps(2) mem_Collect_eq) |
27468 | 137 |
|
75867 | 138 |
lemma injf_max_order_preserving2: |
139 |
assumes "m < n" and E: "\<forall>x. \<exists>y \<in> E. x < y" |
|
140 |
shows "injf_max m E < injf_max n E" |
|
141 |
using \<open>m < n\<close> |
|
142 |
proof (induction n arbitrary: m) |
|
143 |
case 0 then show ?case by auto |
|
144 |
next |
|
145 |
case (Suc n) |
|
146 |
then show ?case |
|
147 |
by (metis E injf_max_order_preserving less_Suc_eq order_less_trans) |
|
148 |
qed |
|
149 |
||
27468 | 150 |
|
64601 | 151 |
lemma inj_injf_max: "\<forall>x. \<exists>y \<in> E. x < y \<Longrightarrow> inj (\<lambda>n. injf_max n E)" |
75867 | 152 |
by (metis injf_max_order_preserving2 linorder_injI order_less_irrefl) |
27468 | 153 |
|
154 |
lemma infinite_set_has_order_preserving_inj: |
|
64601 | 155 |
"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))" |
156 |
for E :: "'a::order set" and f :: "nat \<Rightarrow> 'a" |
|
75867 | 157 |
by (metis image_subsetI inj_injf_max injf_max_mem_set injf_max_order_preserving) |
27468 | 158 |
|
159 |
||
64601 | 160 |
text \<open>Only need the existence of an injective function from \<open>N\<close> to \<open>A\<close> for proof.\<close> |
27468 | 161 |
|
75867 | 162 |
lemma hypnat_infinite_has_nonstandard: |
163 |
assumes "infinite A" |
|
164 |
shows "hypnat_of_nat ` A < ( *s* A)" |
|
165 |
by (metis assms IntE NatStar_hypreal_of_real_Int STAR_star_of_image_subset psubsetI |
|
166 |
infinite_iff_countable_subset inj_fun_not_hypnat_in_SHNat range_subset_mem_starsetNat) |
|
27468 | 167 |
|
64601 | 168 |
lemma starsetNat_eq_hypnat_of_nat_image_finite: "*s* A = hypnat_of_nat ` A \<Longrightarrow> finite A" |
169 |
by (metis hypnat_infinite_has_nonstandard less_irrefl) |
|
170 |
||
171 |
lemma finite_starsetNat_iff: "*s* A = hypnat_of_nat ` A \<longleftrightarrow> finite A" |
|
172 |
by (blast intro!: starsetNat_eq_hypnat_of_nat_image_finite NatStar_hypnat_of_nat) |
|
27468 | 173 |
|
75867 | 174 |
lemma hypnat_infinite_has_nonstandard_iff: "infinite A \<longleftrightarrow> hypnat_of_nat ` A < *s* A" |
175 |
by (metis finite_starsetNat_iff hypnat_infinite_has_nonstandard nless_le) |
|
27468 | 176 |
|
177 |
||
64601 | 178 |
subsection \<open>Existence of Infinitely Many Primes: a Nonstandard Proof\<close> |
27468 | 179 |
|
75867 | 180 |
lemma lemma_not_dvd_hypnat_one [simp]: "\<exists>n \<in> - {0}. \<not> hypnat_of_nat n dvd 1" |
181 |
proof - |
|
182 |
have "\<not> hypnat_of_nat 2 dvd 1" |
|
183 |
by transfer auto |
|
184 |
then show ?thesis |
|
185 |
by (metis ComplI singletonD zero_neq_numeral) |
|
186 |
qed |
|
64601 | 187 |
|
188 |
lemma hypnat_add_one_gt_one: "\<And>N::hypnat. 0 < N \<Longrightarrow> 1 < N + 1" |
|
189 |
by transfer simp |
|
27468 | 190 |
|
55165
f4791db20067
Removed a dependence upon Old_Number_Theory. Simplified a few proofs.
paulson <lp15@cam.ac.uk>
parents:
48488
diff
changeset
|
191 |
lemma hypnat_of_nat_zero_not_prime [simp]: "hypnat_of_nat 0 \<notin> starprime" |
64601 | 192 |
by transfer simp |
27468 | 193 |
|
64601 | 194 |
lemma hypnat_zero_not_prime [simp]: "0 \<notin> starprime" |
195 |
using hypnat_of_nat_zero_not_prime by simp |
|
27468 | 196 |
|
55165
f4791db20067
Removed a dependence upon Old_Number_Theory. Simplified a few proofs.
paulson <lp15@cam.ac.uk>
parents:
48488
diff
changeset
|
197 |
lemma hypnat_of_nat_one_not_prime [simp]: "hypnat_of_nat 1 \<notin> starprime" |
64601 | 198 |
by transfer simp |
27468 | 199 |
|
55165
f4791db20067
Removed a dependence upon Old_Number_Theory. Simplified a few proofs.
paulson <lp15@cam.ac.uk>
parents:
48488
diff
changeset
|
200 |
lemma hypnat_one_not_prime [simp]: "1 \<notin> starprime" |
64601 | 201 |
using hypnat_of_nat_one_not_prime by simp |
27468 | 202 |
|
64601 | 203 |
lemma hdvd_diff: "\<And>k m n :: hypnat. k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m - n)" |
204 |
by transfer (rule dvd_diff_nat) |
|
27468 | 205 |
|
64601 | 206 |
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
|
207 |
by transfer simp |
27468 | 208 |
|
64601 | 209 |
text \<open>Already proved as \<open>primes_infinite\<close>, but now using non-standard naturals.\<close> |
75867 | 210 |
theorem not_finite_prime: "infinite {p::nat. prime p}" |
211 |
proof - |
|
212 |
obtain N k where N: "\<forall>n\<in>- {0}. hypnat_of_nat n dvd N" "k\<in>starprime" "k dvd N + 1" |
|
213 |
by (meson hyperprime_factor_exists hypnat_add_one_gt_one hypnat_dvd_all_hypnat_of_nat) |
|
214 |
then have "k \<noteq> 1" |
|
215 |
using \<open>k \<in> starprime\<close> by force |
|
216 |
then have "k \<notin> hypnat_of_nat ` {p. prime p}" |
|
217 |
using N dvd_add_right_iff hdvd_one_eq_one not_prime_0 by blast |
|
218 |
then show ?thesis |
|
219 |
by (metis \<open>k \<in> starprime\<close> finite_starsetNat_iff starprime_def) |
|
220 |
qed |
|
27468 | 221 |
|
222 |
end |