author | urbanc |
Tue, 13 Dec 2005 18:11:21 +0100 | |
changeset 18396 | b3e7da94b51f |
parent 17429 | e8d6ed3aacfe |
child 19380 | b808efaa5828 |
permissions | -rw-r--r-- |
10751 | 1 |
(* Title : HOL/Real/Hyperreal/HyperDef.thy |
2 |
ID : $Id$ |
|
3 |
Author : Jacques D. Fleuriot |
|
4 |
Copyright : 1998 University of Cambridge |
|
14468 | 5 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2004 |
13487 | 6 |
*) |
10751 | 7 |
|
14468 | 8 |
header{*Construction of Hyperreals Using Ultrafilters*} |
9 |
||
15131 | 10 |
theory HyperDef |
17298 | 11 |
imports StarClasses "../Real/Real" |
16417 | 12 |
uses ("fuf.ML") (*Warning: file fuf.ML refers to the name Hyperdef!*) |
15131 | 13 |
begin |
10751 | 14 |
|
17298 | 15 |
types hypreal = "real star" |
14299 | 16 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
17 |
syntax hypreal_of_real :: "real => real star" |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
18 |
translations "hypreal_of_real" => "star_of :: real => real star" |
10751 | 19 |
|
20 |
constdefs |
|
21 |
||
14691 | 22 |
omega :: hypreal -- {*an infinite number @{text "= [<1,2,3,...>]"} *} |
17298 | 23 |
"omega == star_n (%n. real (Suc n))" |
13487 | 24 |
|
14691 | 25 |
epsilon :: hypreal -- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *} |
17298 | 26 |
"epsilon == star_n (%n. inverse (real (Suc n)))" |
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
27 |
|
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
28 |
syntax (xsymbols) |
14299 | 29 |
omega :: hypreal ("\<omega>") |
30 |
epsilon :: hypreal ("\<epsilon>") |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
31 |
|
14565 | 32 |
syntax (HTML output) |
33 |
omega :: hypreal ("\<omega>") |
|
34 |
epsilon :: hypreal ("\<epsilon>") |
|
35 |
||
14329 | 36 |
|
37 |
subsection{*Existence of Free Ultrafilter over the Naturals*} |
|
38 |
||
39 |
text{*Also, proof of various properties of @{term FreeUltrafilterNat}: |
|
40 |
an arbitrary free ultrafilter*} |
|
14299 | 41 |
|
17290 | 42 |
lemma FreeUltrafilterNat_Ex: "\<exists>U::nat set set. freeultrafilter U" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
43 |
by (rule nat_infinite [THEN freeultrafilter_Ex]) |
14299 | 44 |
|
17298 | 45 |
lemma FreeUltrafilterNat_mem: "freeultrafilter FreeUltrafilterNat" |
14299 | 46 |
apply (unfold FreeUltrafilterNat_def) |
17290 | 47 |
apply (rule someI_ex) |
48 |
apply (rule FreeUltrafilterNat_Ex) |
|
14299 | 49 |
done |
50 |
||
17290 | 51 |
lemma UltrafilterNat_mem: "ultrafilter FreeUltrafilterNat" |
52 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.ultrafilter]) |
|
53 |
||
54 |
lemma FilterNat_mem: "filter FreeUltrafilterNat" |
|
55 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.filter]) |
|
56 |
||
14299 | 57 |
lemma FreeUltrafilterNat_finite: "finite x ==> x \<notin> FreeUltrafilterNat" |
17290 | 58 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.finite]) |
14299 | 59 |
|
14705 | 60 |
lemma FreeUltrafilterNat_not_finite: "x \<in> FreeUltrafilterNat ==> ~ finite x" |
17290 | 61 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.infinite]) |
14299 | 62 |
|
14301 | 63 |
lemma FreeUltrafilterNat_empty [simp]: "{} \<notin> FreeUltrafilterNat" |
17290 | 64 |
by (rule FilterNat_mem [THEN filter.empty]) |
14299 | 65 |
|
14329 | 66 |
lemma FreeUltrafilterNat_Int: |
14705 | 67 |
"[| X \<in> FreeUltrafilterNat; Y \<in> FreeUltrafilterNat |] |
14299 | 68 |
==> X Int Y \<in> FreeUltrafilterNat" |
17290 | 69 |
by (rule FilterNat_mem [THEN filter.Int]) |
14299 | 70 |
|
14329 | 71 |
lemma FreeUltrafilterNat_subset: |
14705 | 72 |
"[| X \<in> FreeUltrafilterNat; X \<subseteq> Y |] |
14299 | 73 |
==> Y \<in> FreeUltrafilterNat" |
17290 | 74 |
by (rule FilterNat_mem [THEN filter.subset]) |
14299 | 75 |
|
14329 | 76 |
lemma FreeUltrafilterNat_Compl: |
14705 | 77 |
"X \<in> FreeUltrafilterNat ==> -X \<notin> FreeUltrafilterNat" |
17290 | 78 |
apply (erule contrapos_pn) |
79 |
apply (erule UltrafilterNat_mem [THEN ultrafilter.not_mem_iff, THEN iffD2]) |
|
80 |
done |
|
14299 | 81 |
|
14329 | 82 |
lemma FreeUltrafilterNat_Compl_mem: |
83 |
"X\<notin> FreeUltrafilterNat ==> -X \<in> FreeUltrafilterNat" |
|
17290 | 84 |
by (rule UltrafilterNat_mem [THEN ultrafilter.not_mem_iff, THEN iffD1]) |
14299 | 85 |
|
14329 | 86 |
lemma FreeUltrafilterNat_Compl_iff1: |
14705 | 87 |
"(X \<notin> FreeUltrafilterNat) = (-X \<in> FreeUltrafilterNat)" |
17290 | 88 |
by (rule UltrafilterNat_mem [THEN ultrafilter.not_mem_iff]) |
14299 | 89 |
|
14329 | 90 |
lemma FreeUltrafilterNat_Compl_iff2: |
14705 | 91 |
"(X \<in> FreeUltrafilterNat) = (-X \<notin> FreeUltrafilterNat)" |
14301 | 92 |
by (auto simp add: FreeUltrafilterNat_Compl_iff1 [symmetric]) |
14299 | 93 |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
94 |
lemma cofinite_mem_FreeUltrafilterNat: "finite (-X) ==> X \<in> FreeUltrafilterNat" |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
95 |
apply (drule FreeUltrafilterNat_finite) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
96 |
apply (simp add: FreeUltrafilterNat_Compl_iff2 [symmetric]) |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
97 |
done |
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
98 |
|
15539 | 99 |
lemma FreeUltrafilterNat_UNIV [iff]: "UNIV \<in> FreeUltrafilterNat" |
17290 | 100 |
by (rule FilterNat_mem [THEN filter.UNIV]) |
14299 | 101 |
|
14329 | 102 |
lemma FreeUltrafilterNat_Nat_set_refl [intro]: |
103 |
"{n. P(n) = P(n)} \<in> FreeUltrafilterNat" |
|
14301 | 104 |
by simp |
14299 | 105 |
|
106 |
lemma FreeUltrafilterNat_P: "{n::nat. P} \<in> FreeUltrafilterNat ==> P" |
|
14301 | 107 |
by (rule ccontr, simp) |
14299 | 108 |
|
109 |
lemma FreeUltrafilterNat_Ex_P: "{n. P(n)} \<in> FreeUltrafilterNat ==> \<exists>n. P(n)" |
|
14301 | 110 |
by (rule ccontr, simp) |
14299 | 111 |
|
112 |
lemma FreeUltrafilterNat_all: "\<forall>n. P(n) ==> {n. P(n)} \<in> FreeUltrafilterNat" |
|
15539 | 113 |
by (auto) |
14299 | 114 |
|
14329 | 115 |
|
116 |
text{*Define and use Ultrafilter tactics*} |
|
14299 | 117 |
use "fuf.ML" |
118 |
||
119 |
method_setup fuf = {* |
|
120 |
Method.ctxt_args (fn ctxt => |
|
121 |
Method.METHOD (fn facts => |
|
15032 | 122 |
fuf_tac (local_clasimpset_of ctxt) 1)) *} |
14299 | 123 |
"free ultrafilter tactic" |
124 |
||
125 |
method_setup ultra = {* |
|
126 |
Method.ctxt_args (fn ctxt => |
|
127 |
Method.METHOD (fn facts => |
|
15032 | 128 |
ultra_tac (local_clasimpset_of ctxt) 1)) *} |
14299 | 129 |
"ultrafilter tactic" |
130 |
||
131 |
||
14329 | 132 |
text{*One further property of our free ultrafilter*} |
133 |
lemma FreeUltrafilterNat_Un: |
|
14705 | 134 |
"X Un Y \<in> FreeUltrafilterNat |
135 |
==> X \<in> FreeUltrafilterNat | Y \<in> FreeUltrafilterNat" |
|
136 |
by (auto, ultra) |
|
14299 | 137 |
|
138 |
||
17298 | 139 |
subsection{*Properties of @{term starrel}*} |
14329 | 140 |
|
17298 | 141 |
text{*Proving that @{term starrel} is an equivalence relation*} |
14299 | 142 |
|
17298 | 143 |
lemma starrel_iff: "((X,Y) \<in> starrel) = ({n. X n = Y n} \<in> FreeUltrafilterNat)" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
144 |
by (rule StarDef.starrel_iff) |
14299 | 145 |
|
17298 | 146 |
lemma starrel_refl: "(x,x) \<in> starrel" |
147 |
by (simp add: starrel_def) |
|
14299 | 148 |
|
17298 | 149 |
lemma starrel_sym [rule_format (no_asm)]: "(x,y) \<in> starrel --> (y,x) \<in> starrel" |
150 |
by (simp add: starrel_def eq_commute) |
|
14299 | 151 |
|
17298 | 152 |
lemma starrel_trans: |
153 |
"[|(x,y) \<in> starrel; (y,z) \<in> starrel|] ==> (x,z) \<in> starrel" |
|
154 |
by (simp add: starrel_def, ultra) |
|
14299 | 155 |
|
17298 | 156 |
lemma equiv_starrel: "equiv UNIV starrel" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
157 |
by (rule StarDef.equiv_starrel) |
14299 | 158 |
|
17298 | 159 |
(* (starrel `` {x} = starrel `` {y}) = ((x,y) \<in> starrel) *) |
160 |
lemmas equiv_starrel_iff = |
|
161 |
eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I, simp] |
|
14299 | 162 |
|
17298 | 163 |
lemma starrel_in_hypreal [simp]: "starrel``{x}:star" |
164 |
by (simp add: star_def starrel_def quotient_def, blast) |
|
14299 | 165 |
|
17298 | 166 |
declare Abs_star_inject [simp] Abs_star_inverse [simp] |
167 |
declare equiv_starrel [THEN eq_equiv_class_iff, simp] |
|
14299 | 168 |
|
17298 | 169 |
lemmas eq_starrelD = eq_equiv_class [OF _ equiv_starrel] |
14299 | 170 |
|
17298 | 171 |
lemma lemma_starrel_refl [simp]: "x \<in> starrel `` {x}" |
172 |
by (simp add: starrel_def) |
|
14299 | 173 |
|
17298 | 174 |
lemma hypreal_empty_not_mem [simp]: "{} \<notin> star" |
175 |
apply (simp add: star_def) |
|
14299 | 176 |
apply (auto elim!: quotientE equalityCE) |
177 |
done |
|
178 |
||
17298 | 179 |
lemma Rep_hypreal_nonempty [simp]: "Rep_star x \<noteq> {}" |
180 |
by (insert Rep_star [of x], auto) |
|
14299 | 181 |
|
14329 | 182 |
subsection{*@{term hypreal_of_real}: |
183 |
the Injection from @{typ real} to @{typ hypreal}*} |
|
14299 | 184 |
|
185 |
lemma inj_hypreal_of_real: "inj(hypreal_of_real)" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
186 |
by (rule inj_onI, simp) |
14299 | 187 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
188 |
lemma Rep_star_star_n_iff [simp]: |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
189 |
"(X \<in> Rep_star (star_n Y)) = ({n. Y n = X n} \<in> \<U>)" |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
190 |
by (simp add: star_n_def) |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
191 |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
192 |
lemma Rep_star_star_n: "X \<in> Rep_star (star_n X)" |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
193 |
by simp |
14329 | 194 |
|
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
195 |
subsection{* Properties of @{term star_n} *} |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
196 |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
197 |
lemma star_n_add: |
17298 | 198 |
"star_n X + star_n Y = star_n (%n. X n + Y n)" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
199 |
by (simp only: star_add_def starfun2_star_n) |
14329 | 200 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
201 |
lemma star_n_minus: |
17298 | 202 |
"- star_n X = star_n (%n. -(X n))" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
203 |
by (simp only: star_minus_def starfun_star_n) |
14299 | 204 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
205 |
lemma star_n_diff: |
17298 | 206 |
"star_n X - star_n Y = star_n (%n. X n - Y n)" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
207 |
by (simp only: star_diff_def starfun2_star_n) |
14299 | 208 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
209 |
lemma star_n_mult: |
17298 | 210 |
"star_n X * star_n Y = star_n (%n. X n * Y n)" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
211 |
by (simp only: star_mult_def starfun2_star_n) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
212 |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
213 |
lemma star_n_inverse: |
17301 | 214 |
"inverse (star_n X) = star_n (%n. inverse(X n))" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
215 |
by (simp only: star_inverse_def starfun_star_n) |
14299 | 216 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
217 |
lemma star_n_le: |
17298 | 218 |
"star_n X \<le> star_n Y = |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
219 |
({n. X n \<le> Y n} \<in> FreeUltrafilterNat)" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
220 |
by (simp only: star_le_def starP2_star_n) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
221 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
222 |
lemma star_n_less: |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
223 |
"star_n X < star_n Y = ({n. X n < Y n} \<in> FreeUltrafilterNat)" |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
224 |
by (simp only: star_less_def starP2_star_n) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
225 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
226 |
lemma star_n_zero_num: "0 = star_n (%n. 0)" |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
227 |
by (simp only: star_zero_def star_of_def) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
228 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
229 |
lemma star_n_one_num: "1 = star_n (%n. 1)" |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
230 |
by (simp only: star_one_def star_of_def) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
231 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
232 |
lemma star_n_abs: |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
233 |
"abs (star_n X) = star_n (%n. abs (X n))" |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
234 |
by (simp only: star_abs_def starfun_star_n) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
235 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
236 |
subsection{*Misc Others*} |
14299 | 237 |
|
14370 | 238 |
lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \<noteq> y" |
15539 | 239 |
by (auto) |
14329 | 240 |
|
14331 | 241 |
lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)" |
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
242 |
by auto |
14331 | 243 |
|
14329 | 244 |
lemma hypreal_mult_left_cancel: "(c::hypreal) \<noteq> 0 ==> (c*a=c*b) = (a=b)" |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
245 |
by auto |
14329 | 246 |
|
247 |
lemma hypreal_mult_right_cancel: "(c::hypreal) \<noteq> 0 ==> (a*c=b*c) = (a=b)" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
248 |
by auto |
14329 | 249 |
|
14301 | 250 |
lemma hypreal_omega_gt_zero [simp]: "0 < omega" |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
251 |
by (simp add: omega_def star_n_zero_num star_n_less) |
14370 | 252 |
|
253 |
subsection{*Existence of Infinite Hyperreal Number*} |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
254 |
|
14370 | 255 |
text{*Existence of infinite number not corresponding to any real number. |
256 |
Use assumption that member @{term FreeUltrafilterNat} is not finite.*} |
|
257 |
||
258 |
||
259 |
text{*A few lemmas first*} |
|
260 |
||
261 |
lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} | |
|
262 |
(\<exists>y. {n::nat. x = real n} = {y})" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
263 |
by force |
14370 | 264 |
|
265 |
lemma lemma_finite_omega_set: "finite {n::nat. x = real n}" |
|
266 |
by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto) |
|
267 |
||
268 |
lemma not_ex_hypreal_of_real_eq_omega: |
|
269 |
"~ (\<exists>x. hypreal_of_real x = omega)" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
270 |
apply (simp add: omega_def) |
17298 | 271 |
apply (simp add: star_of_def star_n_eq_iff) |
14370 | 272 |
apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] |
273 |
lemma_finite_omega_set [THEN FreeUltrafilterNat_finite]) |
|
274 |
done |
|
275 |
||
276 |
lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \<noteq> omega" |
|
14705 | 277 |
by (insert not_ex_hypreal_of_real_eq_omega, auto) |
14370 | 278 |
|
279 |
text{*Existence of infinitesimal number also not corresponding to any |
|
280 |
real number*} |
|
281 |
||
282 |
lemma lemma_epsilon_empty_singleton_disj: |
|
283 |
"{n::nat. x = inverse(real(Suc n))} = {} | |
|
284 |
(\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
285 |
by auto |
14370 | 286 |
|
287 |
lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}" |
|
288 |
by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto) |
|
289 |
||
14705 | 290 |
lemma not_ex_hypreal_of_real_eq_epsilon: "~ (\<exists>x. hypreal_of_real x = epsilon)" |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
291 |
by (auto simp add: epsilon_def star_of_def star_n_eq_iff |
14705 | 292 |
lemma_finite_epsilon_set [THEN FreeUltrafilterNat_finite]) |
14370 | 293 |
|
294 |
lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> epsilon" |
|
14705 | 295 |
by (insert not_ex_hypreal_of_real_eq_epsilon, auto) |
14370 | 296 |
|
297 |
lemma hypreal_epsilon_not_zero: "epsilon \<noteq> 0" |
|
17298 | 298 |
by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff |
299 |
del: star_of_zero) |
|
14370 | 300 |
|
301 |
lemma hypreal_epsilon_inverse_omega: "epsilon = inverse(omega)" |
|
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
302 |
by (simp add: epsilon_def omega_def star_n_inverse) |
14370 | 303 |
|
304 |
||
14299 | 305 |
ML |
306 |
{* |
|
307 |
val omega_def = thm "omega_def"; |
|
308 |
val epsilon_def = thm "epsilon_def"; |
|
309 |
||
310 |
val FreeUltrafilterNat_Ex = thm "FreeUltrafilterNat_Ex"; |
|
311 |
val FreeUltrafilterNat_mem = thm "FreeUltrafilterNat_mem"; |
|
312 |
val FreeUltrafilterNat_finite = thm "FreeUltrafilterNat_finite"; |
|
313 |
val FreeUltrafilterNat_not_finite = thm "FreeUltrafilterNat_not_finite"; |
|
314 |
val FreeUltrafilterNat_empty = thm "FreeUltrafilterNat_empty"; |
|
315 |
val FreeUltrafilterNat_Int = thm "FreeUltrafilterNat_Int"; |
|
316 |
val FreeUltrafilterNat_subset = thm "FreeUltrafilterNat_subset"; |
|
317 |
val FreeUltrafilterNat_Compl = thm "FreeUltrafilterNat_Compl"; |
|
318 |
val FreeUltrafilterNat_Compl_mem = thm "FreeUltrafilterNat_Compl_mem"; |
|
319 |
val FreeUltrafilterNat_Compl_iff1 = thm "FreeUltrafilterNat_Compl_iff1"; |
|
320 |
val FreeUltrafilterNat_Compl_iff2 = thm "FreeUltrafilterNat_Compl_iff2"; |
|
321 |
val FreeUltrafilterNat_UNIV = thm "FreeUltrafilterNat_UNIV"; |
|
322 |
val FreeUltrafilterNat_Nat_set_refl = thm "FreeUltrafilterNat_Nat_set_refl"; |
|
323 |
val FreeUltrafilterNat_P = thm "FreeUltrafilterNat_P"; |
|
324 |
val FreeUltrafilterNat_Ex_P = thm "FreeUltrafilterNat_Ex_P"; |
|
325 |
val FreeUltrafilterNat_all = thm "FreeUltrafilterNat_all"; |
|
326 |
val FreeUltrafilterNat_Un = thm "FreeUltrafilterNat_Un"; |
|
17298 | 327 |
val starrel_iff = thm "starrel_iff"; |
328 |
val starrel_in_hypreal = thm "starrel_in_hypreal"; |
|
329 |
val Abs_star_inverse = thm "Abs_star_inverse"; |
|
330 |
val lemma_starrel_refl = thm "lemma_starrel_refl"; |
|
14299 | 331 |
val hypreal_empty_not_mem = thm "hypreal_empty_not_mem"; |
332 |
val Rep_hypreal_nonempty = thm "Rep_hypreal_nonempty"; |
|
333 |
val inj_hypreal_of_real = thm "inj_hypreal_of_real"; |
|
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
334 |
(* val eq_Abs_star = thm "eq_Abs_star"; *) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
335 |
val star_n_minus = thm "star_n_minus"; |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
336 |
val star_n_add = thm "star_n_add"; |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
337 |
val star_n_diff = thm "star_n_diff"; |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
338 |
val star_n_mult = thm "star_n_mult"; |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
339 |
val star_n_inverse = thm "star_n_inverse"; |
14299 | 340 |
val hypreal_mult_left_cancel = thm "hypreal_mult_left_cancel"; |
341 |
val hypreal_mult_right_cancel = thm "hypreal_mult_right_cancel"; |
|
342 |
val hypreal_not_refl2 = thm "hypreal_not_refl2"; |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
343 |
val star_n_less = thm "star_n_less"; |
14299 | 344 |
val hypreal_eq_minus_iff = thm "hypreal_eq_minus_iff"; |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
345 |
val star_n_le = thm "star_n_le"; |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
346 |
val star_n_zero_num = thm "star_n_zero_num"; |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
347 |
val star_n_one_num = thm "star_n_one_num"; |
14299 | 348 |
val hypreal_omega_gt_zero = thm "hypreal_omega_gt_zero"; |
14370 | 349 |
|
350 |
val lemma_omega_empty_singleton_disj = thm"lemma_omega_empty_singleton_disj"; |
|
351 |
val lemma_finite_omega_set = thm"lemma_finite_omega_set"; |
|
352 |
val not_ex_hypreal_of_real_eq_omega = thm"not_ex_hypreal_of_real_eq_omega"; |
|
353 |
val hypreal_of_real_not_eq_omega = thm"hypreal_of_real_not_eq_omega"; |
|
354 |
val not_ex_hypreal_of_real_eq_epsilon = thm"not_ex_hypreal_of_real_eq_epsilon"; |
|
355 |
val hypreal_of_real_not_eq_epsilon = thm"hypreal_of_real_not_eq_epsilon"; |
|
356 |
val hypreal_epsilon_not_zero = thm"hypreal_epsilon_not_zero"; |
|
357 |
val hypreal_epsilon_inverse_omega = thm"hypreal_epsilon_inverse_omega"; |
|
14299 | 358 |
*} |
359 |
||
10751 | 360 |
end |