| author | huffman | 
| Wed, 23 Aug 2006 21:57:43 +0200 | |
| changeset 20407 | 93a34d5d1dc5 | 
| parent 19736 | d8d0f8f51d69 | 
| child 20733 | 4ccef1ac4c9b | 
| permissions | -rw-r--r-- | 
| 13957 | 1 | (* Title : NSPrimes.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 2002 University of Edinburgh | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 4 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 13957 | 5 | *) | 
| 6 | ||
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 7 | header{*The Nonstandard Primes as an Extension of the Prime Numbers*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 8 | |
| 15149 | 9 | theory NSPrimes | 
| 10 | imports "~~/src/HOL/NumberTheory/Factorization" Complex_Main | |
| 11 | begin | |
| 13957 | 12 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 13 | text{*These can be used to derive an alternative proof of the infinitude of
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 14 | primes by considering a property of nonstandard sets.*} | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 15 | |
| 19736 | 16 | definition | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 17 | hdvd :: "[hypnat, hypnat] => bool" (infixl "hdvd" 50) | 
| 19736 | 18 | [transfer_unfold]: "(M::hypnat) hdvd N = ( *p2* (op dvd)) M N" | 
| 17331 | 19 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 20 | starprime :: "hypnat set" | 
| 19736 | 21 |   [transfer_unfold]: "starprime = ( *s* {p. prime p})"
 | 
| 13957 | 22 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 23 | choicefun :: "'a set => 'a" | 
| 19736 | 24 |   "choicefun E = (@x. \<exists>X \<in> Pow(E) -{{}}. x : X)"
 | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 25 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 26 | consts injf_max :: "nat => ('a::{order} set) => 'a"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 27 | primrec | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 28 | injf_max_zero: "injf_max 0 E = choicefun E" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 29 |   injf_max_Suc:  "injf_max (Suc n) E = choicefun({e. e:E & injf_max n E < e})"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 30 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 31 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 32 | text{*A "choice" theorem for ultrafilters, like almost everywhere
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 33 | quantification*} | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 34 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 35 | lemma UF_choice: "{n. \<exists>m. Q n m} : FreeUltrafilterNat
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 36 |       ==> \<exists>f. {n. Q n (f n)} : FreeUltrafilterNat"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 37 | apply (rule_tac x = "%n. (@x. Q n x) " in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 38 | apply (ultra, rule someI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 39 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 40 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 41 | lemma UF_if: "({n. P n} : FreeUltrafilterNat --> {n. Q n} : FreeUltrafilterNat) =
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 42 |       ({n. P n --> Q n} : FreeUltrafilterNat)"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 43 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 44 | apply ultra+ | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 45 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 46 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 47 | lemma UF_conj: "({n. P n} : FreeUltrafilterNat & {n. Q n} : FreeUltrafilterNat) =
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 48 |       ({n. P n & Q n} : FreeUltrafilterNat)"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 49 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 50 | apply ultra+ | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 51 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 52 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 53 | lemma UF_choice_ccontr: "(\<forall>f. {n. Q n (f n)} : FreeUltrafilterNat) =
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 54 |       ({n. \<forall>m. Q n m} : FreeUltrafilterNat)"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 55 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 56 | prefer 2 apply ultra | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 57 | apply (rule ccontr) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 58 | apply (rule contrapos_np) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 59 | apply (erule_tac [2] asm_rl) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 60 | apply (simp (no_asm) add: FreeUltrafilterNat_Compl_iff1 Collect_neg_eq [symmetric]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 61 | apply (rule UF_choice, ultra) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 62 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 63 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 64 | lemma dvd_by_all: "\<forall>M. \<exists>N. 0 < N & (\<forall>m. 0 < m & (m::nat) <= M --> m dvd N)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 65 | apply (rule allI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 66 | apply (induct_tac "M", auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 67 | apply (rule_tac x = "N * (Suc n) " in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 68 | apply (safe, force) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 69 | apply (drule le_imp_less_or_eq, erule disjE) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 70 | apply (force intro!: dvd_mult2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 71 | apply (force intro!: dvd_mult) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 72 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 73 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 74 | lemmas dvd_by_all2 = dvd_by_all [THEN spec, standard] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 75 | |
| 17319 | 76 | lemma lemma_hypnat_P_EX: "(\<exists>(x::hypnat). P x) = (\<exists>f. P (star_n f))" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 77 | apply auto | 
| 17319 | 78 | apply (rule_tac x = x in star_cases, auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 79 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 80 | |
| 17319 | 81 | lemma lemma_hypnat_P_ALL: "(\<forall>(x::hypnat). P x) = (\<forall>f. P (star_n f))" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 82 | apply auto | 
| 17319 | 83 | apply (rule_tac x = x in star_cases, auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 84 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 85 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 86 | lemma hdvd: | 
| 17319 | 87 | "(star_n X hdvd star_n Y) = | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 88 |        ({n. X n dvd Y n} : FreeUltrafilterNat)"
 | 
| 17319 | 89 | by (simp add: hdvd_def starP2) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 90 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 91 | lemma hypnat_of_nat_le_zero_iff: "(hypnat_of_nat n <= 0) = (n = 0)" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17331diff
changeset | 92 | by (transfer, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 93 | declare hypnat_of_nat_le_zero_iff [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 94 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 95 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 96 | (* Goldblatt: Exercise 5.11(2) - p. 57 *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 97 | lemma hdvd_by_all: "\<forall>M. \<exists>N. 0 < N & (\<forall>m. 0 < m & (m::hypnat) <= M --> m hdvd N)" | 
| 17331 | 98 | by (transfer, rule dvd_by_all) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 99 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 100 | lemmas hdvd_by_all2 = hdvd_by_all [THEN spec, standard] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 101 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 102 | (* Goldblatt: Exercise 5.11(2) - p. 57 *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 103 | lemma hypnat_dvd_all_hypnat_of_nat: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 104 |      "\<exists>(N::hypnat). 0 < N & (\<forall>n \<in> -{0::nat}. hypnat_of_nat(n) hdvd N)"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 105 | apply (cut_tac hdvd_by_all) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 106 | apply (drule_tac x = whn in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 107 | apply (rule exI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 108 | apply (drule_tac x = "hypnat_of_nat n" in spec) | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17331diff
changeset | 109 | apply (auto simp add: linorder_not_less star_of_eq_0) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 110 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 111 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 112 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 113 | text{*The nonstandard extension of the set prime numbers consists of precisely
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 114 | those hypernaturals exceeding 1 that have no nontrivial factors*} | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 115 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 116 | (* Goldblatt: Exercise 5.11(3a) - p 57 *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 117 | lemma starprime: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 118 |   "starprime = {p. 1 < p & (\<forall>m. m hdvd p --> m = 1 | m = p)}"
 | 
| 17331 | 119 | by (transfer, auto simp add: prime_def) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 120 | |
| 16663 | 121 | lemma prime_two: "prime 2" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 122 | apply (unfold prime_def, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 123 | apply (frule dvd_imp_le) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 124 | apply (auto dest: dvd_0_left) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 125 | apply (case_tac m, simp, arith) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 126 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 127 | declare prime_two [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 128 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 129 | (* proof uses course-of-value induction *) | 
| 16663 | 130 | lemma prime_factor_exists [rule_format]: "Suc 0 < n --> (\<exists>k. prime k & k dvd n)" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 131 | apply (rule_tac n = n in nat_less_induct, auto) | 
| 16663 | 132 | apply (case_tac "prime n") | 
| 133 | apply (rule_tac x = n in exI, auto) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 134 | apply (drule conjI [THEN not_prime_ex_mk], auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 135 | apply (drule_tac x = m in spec, auto) | 
| 16663 | 136 | apply (rule_tac x = ka in exI) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 137 | apply (auto intro: dvd_mult2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 138 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 139 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 140 | (* Goldblatt Exercise 5.11(3b) - p 57 *) | 
| 17319 | 141 | lemma hyperprime_factor_exists [rule_format]: | 
| 142 | "!!n. 1 < n ==> (\<exists>k \<in> starprime. k hdvd n)" | |
| 17331 | 143 | by (transfer, simp add: prime_factor_exists) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 144 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 145 | (* Goldblatt Exercise 3.10(1) - p. 29 *) | 
| 17319 | 146 | lemma NatStar_hypnat_of_nat: "finite A ==> *s* A = hypnat_of_nat ` A" | 
| 147 | apply (rule_tac P = "%x. *s* x = hypnat_of_nat ` x" in finite_induct) | |
| 148 | apply auto | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 149 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 150 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 151 | (* proved elsewhere? *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 152 | lemma FreeUltrafilterNat_singleton_not_mem: "{x} \<notin> FreeUltrafilterNat"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 153 | by (auto intro!: FreeUltrafilterNat_finite) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 154 | declare FreeUltrafilterNat_singleton_not_mem [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 155 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 156 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 157 | subsection{*Another characterization of infinite set of natural numbers*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 158 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 159 | lemma finite_nat_set_bounded: "finite N ==> \<exists>n. (\<forall>i \<in> N. i<(n::nat))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 160 | apply (erule_tac F = N in finite_induct, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 161 | apply (rule_tac x = "Suc n + x" in exI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 162 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 163 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 164 | lemma finite_nat_set_bounded_iff: "finite N = (\<exists>n. (\<forall>i \<in> N. i<(n::nat)))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 165 | by (blast intro: finite_nat_set_bounded bounded_nat_set_is_finite) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 166 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 167 | lemma not_finite_nat_set_iff: "(~ finite N) = (\<forall>n. \<exists>i \<in> N. n <= (i::nat))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 168 | by (auto simp add: finite_nat_set_bounded_iff le_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 169 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 170 | lemma bounded_nat_set_is_finite2: "(\<forall>i \<in> N. i<=(n::nat)) ==> finite N" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 171 | apply (rule finite_subset) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 172 | apply (rule_tac [2] finite_atMost, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 173 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 174 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 175 | lemma finite_nat_set_bounded2: "finite N ==> \<exists>n. (\<forall>i \<in> N. i<=(n::nat))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 176 | apply (erule_tac F = N in finite_induct, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 177 | apply (rule_tac x = "n + x" in exI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 178 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 179 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 180 | lemma finite_nat_set_bounded_iff2: "finite N = (\<exists>n. (\<forall>i \<in> N. i<=(n::nat)))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 181 | by (blast intro: finite_nat_set_bounded2 bounded_nat_set_is_finite2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 182 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 183 | lemma not_finite_nat_set_iff2: "(~ finite N) = (\<forall>n. \<exists>i \<in> N. n < (i::nat))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 184 | by (auto simp add: finite_nat_set_bounded_iff2 le_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 185 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 186 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 187 | subsection{*An injective function cannot define an embedded natural number*}
 | 
| 13957 | 188 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 189 | lemma lemma_infinite_set_singleton: "\<forall>m n. m \<noteq> n --> f n \<noteq> f m | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 190 |       ==>  {n. f n = N} = {} |  (\<exists>m. {n. f n = N} = {m})"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 191 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 192 | apply (drule_tac x = x in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 193 | apply (subgoal_tac "\<forall>n. (f n = f x) = (x = n) ") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 194 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 195 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 196 | |
| 17319 | 197 | lemma inj_fun_not_hypnat_in_SHNat: "inj (f::nat=>nat) ==> star_n f \<notin> Nats" | 
| 198 | apply (auto simp add: SHNat_eq hypnat_of_nat_eq star_n_eq_iff) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 199 | apply (subgoal_tac "\<forall>m n. m \<noteq> n --> f n \<noteq> f m", auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 200 | apply (drule_tac [2] injD) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 201 | prefer 2 apply assumption | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 202 | apply (drule_tac N = N in lemma_infinite_set_singleton, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 203 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 204 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 205 | lemma range_subset_mem_starsetNat: | 
| 17319 | 206 | "range f <= A ==> star_n f \<in> *s* A" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17331diff
changeset | 207 | apply (simp add: starset_def star_of_def Iset_star_n) | 
| 17319 | 208 | apply (subgoal_tac "\<forall>n. f n \<in> A", auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 209 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 210 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 211 | (*--------------------------------------------------------------------------------*) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 212 | (* Gleason Proposition 11-5.5. pg 149, pg 155 (ex. 3) and pg. 360 *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 213 | (* Let E be a nonvoid ordered set with no maximal elements (note: effectively an *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 214 | (* infinite set if we take E = N (Nats)). Then there exists an order-preserving *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 215 | (* injection from N to E. Of course, (as some doofus will undoubtedly point out! *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 216 | (* :-)) can use notion of least element in proof (i.e. no need for choice) if *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 217 | (* dealing with nats as we have well-ordering property *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 218 | (*--------------------------------------------------------------------------------*) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 219 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 220 | lemma lemmaPow3: "E \<noteq> {} ==> \<exists>x. \<exists>X \<in> (Pow E - {{}}). x: X"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 221 | by auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 222 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 223 | lemma choicefun_mem_set: "E \<noteq> {} ==> choicefun E \<in> E"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 224 | apply (unfold choicefun_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 225 | apply (rule lemmaPow3 [THEN someI2_ex], auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 226 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 227 | declare choicefun_mem_set [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 228 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 229 | lemma injf_max_mem_set: "[| E \<noteq>{}; \<forall>x. \<exists>y \<in> E. x < y |] ==> injf_max n E \<in> E"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 230 | apply (induct_tac "n", force) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 231 | apply (simp (no_asm) add: choicefun_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 232 | apply (rule lemmaPow3 [THEN someI2_ex], auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 233 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 234 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 235 | lemma injf_max_order_preserving: "\<forall>x. \<exists>y \<in> E. x < y ==> injf_max n E < injf_max (Suc n) E" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 236 | apply (simp (no_asm) add: choicefun_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 237 | apply (rule lemmaPow3 [THEN someI2_ex], auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 238 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 239 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 240 | lemma injf_max_order_preserving2: "\<forall>x. \<exists>y \<in> E. x < y | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 241 | ==> \<forall>n m. m < n --> injf_max m E < injf_max n E" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 242 | apply (rule allI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 243 | apply (induct_tac "n", auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 244 | apply (simp (no_asm) add: choicefun_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 245 | apply (rule lemmaPow3 [THEN someI2_ex]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 246 | apply (auto simp add: less_Suc_eq) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 247 | apply (drule_tac x = m in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 248 | apply (drule subsetD, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 249 | apply (drule_tac x = "injf_max m E" in order_less_trans, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 250 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 251 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 252 | lemma inj_injf_max: "\<forall>x. \<exists>y \<in> E. x < y ==> inj (%n. injf_max n E)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 253 | apply (rule inj_onI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 254 | apply (rule ccontr, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 255 | apply (drule injf_max_order_preserving2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 256 | apply (cut_tac m = x and n = y in less_linear, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 257 | apply (auto dest!: spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 258 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 259 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 260 | lemma infinite_set_has_order_preserving_inj: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 261 |      "[| (E::('a::{order} set)) \<noteq> {}; \<forall>x. \<exists>y \<in> E. x < y |]
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 262 | ==> \<exists>f. range f <= E & inj (f::nat => 'a) & (\<forall>m. f m < f(Suc m))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 263 | apply (rule_tac x = "%n. injf_max n E" in exI, safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 264 | apply (rule injf_max_mem_set) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 265 | apply (rule_tac [3] inj_injf_max) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 266 | apply (rule_tac [4] injf_max_order_preserving, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 267 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 268 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 269 | text{*Only need the existence of an injective function from N to A for proof*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 270 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 271 | lemma hypnat_infinite_has_nonstandard: | 
| 17319 | 272 | "~ finite A ==> hypnat_of_nat ` A < ( *s* A)" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 273 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 274 | apply (subgoal_tac "A \<noteq> {}")
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 275 | prefer 2 apply force | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 276 | apply (drule infinite_set_has_order_preserving_inj) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 277 | apply (erule not_finite_nat_set_iff2 [THEN iffD1], auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 278 | apply (drule inj_fun_not_hypnat_in_SHNat) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 279 | apply (drule range_subset_mem_starsetNat) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 280 | apply (auto simp add: SHNat_eq) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 281 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 282 | |
| 17319 | 283 | lemma starsetNat_eq_hypnat_of_nat_image_finite: "*s* A = hypnat_of_nat ` A ==> finite A" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 284 | apply (rule ccontr) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 285 | apply (auto dest: hypnat_infinite_has_nonstandard) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 286 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 287 | |
| 17319 | 288 | lemma finite_starsetNat_iff: "( *s* A = hypnat_of_nat ` A) = (finite A)" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 289 | by (blast intro!: starsetNat_eq_hypnat_of_nat_image_finite NatStar_hypnat_of_nat) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 290 | |
| 17319 | 291 | lemma hypnat_infinite_has_nonstandard_iff: "(~ finite A) = (hypnat_of_nat ` A < *s* A)" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 292 | apply (rule iffI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 293 | apply (blast intro!: hypnat_infinite_has_nonstandard) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 294 | apply (auto simp add: finite_starsetNat_iff [symmetric]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 295 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 296 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 297 | subsection{*Existence of Infinitely Many Primes: a Nonstandard Proof*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 298 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 299 | lemma lemma_not_dvd_hypnat_one: "~ (\<forall>n \<in> - {0}. hypnat_of_nat n hdvd 1)"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 300 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 301 | apply (rule_tac x = 2 in bexI) | 
| 17331 | 302 | apply (transfer, auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 303 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 304 | declare lemma_not_dvd_hypnat_one [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 305 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 306 | lemma lemma_not_dvd_hypnat_one2: "\<exists>n \<in> - {0}. ~ hypnat_of_nat n hdvd 1"
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 307 | apply (cut_tac lemma_not_dvd_hypnat_one) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 308 | apply (auto simp del: lemma_not_dvd_hypnat_one) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 309 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 310 | declare lemma_not_dvd_hypnat_one2 [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 311 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 312 | (* not needed here *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 313 | lemma hypnat_gt_zero_gt_one: | 
| 17319 | 314 | "!!N. [| 0 < (N::hypnat); N \<noteq> 1 |] ==> 1 < N" | 
| 315 | by (transfer, simp) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 316 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 317 | lemma hypnat_add_one_gt_one: | 
| 17319 | 318 | "!!N. 0 < N ==> 1 < (N::hypnat) + 1" | 
| 319 | by (transfer, simp) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 320 | |
| 16663 | 321 | lemma zero_not_prime: "\<not> prime 0" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 322 | apply safe | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 323 | apply (drule prime_g_zero, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 324 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 325 | declare zero_not_prime [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 326 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 327 | lemma hypnat_of_nat_zero_not_prime: "hypnat_of_nat 0 \<notin> starprime" | 
| 17331 | 328 | by (transfer, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 329 | declare hypnat_of_nat_zero_not_prime [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 330 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 331 | lemma hypnat_zero_not_prime: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 332 | "0 \<notin> starprime" | 
| 17319 | 333 | by (cut_tac hypnat_of_nat_zero_not_prime, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 334 | declare hypnat_zero_not_prime [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 335 | |
| 16663 | 336 | lemma one_not_prime: "\<not> prime 1" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 337 | apply safe | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 338 | apply (drule prime_g_one, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 339 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 340 | declare one_not_prime [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 341 | |
| 16663 | 342 | lemma one_not_prime2: "\<not> prime(Suc 0)" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 343 | apply safe | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 344 | apply (drule prime_g_one, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 345 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 346 | declare one_not_prime2 [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 347 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 348 | lemma hypnat_of_nat_one_not_prime: "hypnat_of_nat 1 \<notin> starprime" | 
| 17331 | 349 | by (transfer, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 350 | declare hypnat_of_nat_one_not_prime [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 351 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 352 | lemma hypnat_one_not_prime: "1 \<notin> starprime" | 
| 17319 | 353 | by (cut_tac hypnat_of_nat_one_not_prime, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 354 | declare hypnat_one_not_prime [simp] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 355 | |
| 17319 | 356 | lemma hdvd_diff: "!!k m n. [| k hdvd m; k hdvd n |] ==> k hdvd (m - n)" | 
| 17331 | 357 | by (transfer, rule dvd_diff) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 358 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 359 | lemma dvd_one_eq_one: "x dvd (1::nat) ==> x = 1" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 360 | by (unfold dvd_def, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 361 | |
| 17319 | 362 | lemma hdvd_one_eq_one: "!!x. x hdvd 1 ==> x = 1" | 
| 17331 | 363 | by (transfer, rule dvd_one_eq_one) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 364 | |
| 16663 | 365 | theorem not_finite_prime: "~ finite {p. prime p}"
 | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 366 | apply (rule hypnat_infinite_has_nonstandard_iff [THEN iffD2]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 367 | apply (cut_tac hypnat_dvd_all_hypnat_of_nat) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 368 | apply (erule exE) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 369 | apply (erule conjE) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 370 | apply (subgoal_tac "1 < N + 1") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 371 | prefer 2 apply (blast intro: hypnat_add_one_gt_one) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 372 | apply (drule hyperprime_factor_exists) | 
| 17319 | 373 | apply (auto intro: STAR_mem) | 
| 16663 | 374 | apply (subgoal_tac "k \<notin> hypnat_of_nat ` {p. prime p}")
 | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 375 | apply (force simp add: starprime_def, safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 376 | apply (drule_tac x = x in bspec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 377 | apply (rule ccontr, simp) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 378 | apply (drule hdvd_diff, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 379 | apply (auto dest: hdvd_one_eq_one) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
14051diff
changeset | 380 | done | 
| 13957 | 381 | |
| 382 | end |