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