| author | huffman |
| Fri, 09 Sep 2005 19:34:22 +0200 | |
| changeset 17318 | bc1c75855f3d |
| parent 17301 | 6c82a5c10076 |
| child 17429 | e8d6ed3aacfe |
| 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 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
20 |
typed_print_translation {*
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
21 |
let |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
22 |
fun hr_tr' _ (Type ("fun", (Type ("real", []) :: _))) [t] =
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
23 |
Syntax.const "hypreal_of_real" $ t |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
24 |
| hr_tr' _ _ _ = raise Match; |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
25 |
in [("star_of", hr_tr')] end
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
26 |
*} |
| 13487 | 27 |
|
| 10751 | 28 |
constdefs |
29 |
||
| 14691 | 30 |
omega :: hypreal -- {*an infinite number @{text "= [<1,2,3,...>]"} *}
|
| 17298 | 31 |
"omega == star_n (%n. real (Suc n))" |
| 13487 | 32 |
|
| 14691 | 33 |
epsilon :: hypreal -- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *}
|
| 17298 | 34 |
"epsilon == star_n (%n. inverse (real (Suc n)))" |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
35 |
|
|
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
36 |
syntax (xsymbols) |
| 14299 | 37 |
omega :: hypreal ("\<omega>")
|
38 |
epsilon :: hypreal ("\<epsilon>")
|
|
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
39 |
|
| 14565 | 40 |
syntax (HTML output) |
41 |
omega :: hypreal ("\<omega>")
|
|
42 |
epsilon :: hypreal ("\<epsilon>")
|
|
43 |
||
| 14329 | 44 |
|
45 |
subsection{*The Set of Naturals is not Finite*}
|
|
| 14299 | 46 |
|
47 |
(*** based on James' proof that the set of naturals is not finite ***) |
|
| 14329 | 48 |
lemma finite_exhausts [rule_format]: |
49 |
"finite (A::nat set) --> (\<exists>n. \<forall>m. Suc (n + m) \<notin> A)" |
|
| 14299 | 50 |
apply (rule impI) |
| 14301 | 51 |
apply (erule_tac F = A in finite_induct) |
52 |
apply (blast, erule exE) |
|
| 14299 | 53 |
apply (rule_tac x = "n + x" in exI) |
| 14301 | 54 |
apply (rule allI, erule_tac x = "x + m" in allE) |
| 14299 | 55 |
apply (auto simp add: add_ac) |
56 |
done |
|
57 |
||
| 14329 | 58 |
lemma finite_not_covers [rule_format (no_asm)]: |
59 |
"finite (A :: nat set) --> (\<exists>n. n \<notin>A)" |
|
| 14301 | 60 |
by (rule impI, drule finite_exhausts, blast) |
| 14299 | 61 |
|
62 |
lemma not_finite_nat: "~ finite(UNIV:: nat set)" |
|
| 14301 | 63 |
by (fast dest!: finite_exhausts) |
| 14299 | 64 |
|
| 14329 | 65 |
|
66 |
subsection{*Existence of Free Ultrafilter over the Naturals*}
|
|
67 |
||
68 |
text{*Also, proof of various properties of @{term FreeUltrafilterNat}:
|
|
69 |
an arbitrary free ultrafilter*} |
|
| 14299 | 70 |
|
| 17290 | 71 |
lemma FreeUltrafilterNat_Ex: "\<exists>U::nat set set. freeultrafilter U" |
72 |
by (rule not_finite_nat [THEN freeultrafilter_Ex]) |
|
| 14299 | 73 |
|
| 17298 | 74 |
lemma FreeUltrafilterNat_mem: "freeultrafilter FreeUltrafilterNat" |
| 14299 | 75 |
apply (unfold FreeUltrafilterNat_def) |
| 17290 | 76 |
apply (rule someI_ex) |
77 |
apply (rule FreeUltrafilterNat_Ex) |
|
| 14299 | 78 |
done |
79 |
||
| 17290 | 80 |
lemma UltrafilterNat_mem: "ultrafilter FreeUltrafilterNat" |
81 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.ultrafilter]) |
|
82 |
||
83 |
lemma FilterNat_mem: "filter FreeUltrafilterNat" |
|
84 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.filter]) |
|
85 |
||
| 14299 | 86 |
lemma FreeUltrafilterNat_finite: "finite x ==> x \<notin> FreeUltrafilterNat" |
| 17290 | 87 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.finite]) |
| 14299 | 88 |
|
| 14705 | 89 |
lemma FreeUltrafilterNat_not_finite: "x \<in> FreeUltrafilterNat ==> ~ finite x" |
| 17290 | 90 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.infinite]) |
| 14299 | 91 |
|
| 14301 | 92 |
lemma FreeUltrafilterNat_empty [simp]: "{} \<notin> FreeUltrafilterNat"
|
| 17290 | 93 |
by (rule FilterNat_mem [THEN filter.empty]) |
| 14299 | 94 |
|
| 14329 | 95 |
lemma FreeUltrafilterNat_Int: |
| 14705 | 96 |
"[| X \<in> FreeUltrafilterNat; Y \<in> FreeUltrafilterNat |] |
| 14299 | 97 |
==> X Int Y \<in> FreeUltrafilterNat" |
| 17290 | 98 |
by (rule FilterNat_mem [THEN filter.Int]) |
| 14299 | 99 |
|
| 14329 | 100 |
lemma FreeUltrafilterNat_subset: |
| 14705 | 101 |
"[| X \<in> FreeUltrafilterNat; X \<subseteq> Y |] |
| 14299 | 102 |
==> Y \<in> FreeUltrafilterNat" |
| 17290 | 103 |
by (rule FilterNat_mem [THEN filter.subset]) |
| 14299 | 104 |
|
| 14329 | 105 |
lemma FreeUltrafilterNat_Compl: |
| 14705 | 106 |
"X \<in> FreeUltrafilterNat ==> -X \<notin> FreeUltrafilterNat" |
| 17290 | 107 |
apply (erule contrapos_pn) |
108 |
apply (erule UltrafilterNat_mem [THEN ultrafilter.not_mem_iff, THEN iffD2]) |
|
109 |
done |
|
| 14299 | 110 |
|
| 14329 | 111 |
lemma FreeUltrafilterNat_Compl_mem: |
112 |
"X\<notin> FreeUltrafilterNat ==> -X \<in> FreeUltrafilterNat" |
|
| 17290 | 113 |
by (rule UltrafilterNat_mem [THEN ultrafilter.not_mem_iff, THEN iffD1]) |
| 14299 | 114 |
|
| 14329 | 115 |
lemma FreeUltrafilterNat_Compl_iff1: |
| 14705 | 116 |
"(X \<notin> FreeUltrafilterNat) = (-X \<in> FreeUltrafilterNat)" |
| 17290 | 117 |
by (rule UltrafilterNat_mem [THEN ultrafilter.not_mem_iff]) |
| 14299 | 118 |
|
| 14329 | 119 |
lemma FreeUltrafilterNat_Compl_iff2: |
| 14705 | 120 |
"(X \<in> FreeUltrafilterNat) = (-X \<notin> FreeUltrafilterNat)" |
| 14301 | 121 |
by (auto simp add: FreeUltrafilterNat_Compl_iff1 [symmetric]) |
| 14299 | 122 |
|
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
123 |
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
|
124 |
apply (drule FreeUltrafilterNat_finite) |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
125 |
apply (simp add: FreeUltrafilterNat_Compl_iff2 [symmetric]) |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
126 |
done |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
127 |
|
| 15539 | 128 |
lemma FreeUltrafilterNat_UNIV [iff]: "UNIV \<in> FreeUltrafilterNat" |
| 17290 | 129 |
by (rule FilterNat_mem [THEN filter.UNIV]) |
| 14299 | 130 |
|
| 14329 | 131 |
lemma FreeUltrafilterNat_Nat_set_refl [intro]: |
132 |
"{n. P(n) = P(n)} \<in> FreeUltrafilterNat"
|
|
| 14301 | 133 |
by simp |
| 14299 | 134 |
|
135 |
lemma FreeUltrafilterNat_P: "{n::nat. P} \<in> FreeUltrafilterNat ==> P"
|
|
| 14301 | 136 |
by (rule ccontr, simp) |
| 14299 | 137 |
|
138 |
lemma FreeUltrafilterNat_Ex_P: "{n. P(n)} \<in> FreeUltrafilterNat ==> \<exists>n. P(n)"
|
|
| 14301 | 139 |
by (rule ccontr, simp) |
| 14299 | 140 |
|
141 |
lemma FreeUltrafilterNat_all: "\<forall>n. P(n) ==> {n. P(n)} \<in> FreeUltrafilterNat"
|
|
| 15539 | 142 |
by (auto) |
| 14299 | 143 |
|
| 14329 | 144 |
|
145 |
text{*Define and use Ultrafilter tactics*}
|
|
| 14299 | 146 |
use "fuf.ML" |
147 |
||
148 |
method_setup fuf = {*
|
|
149 |
Method.ctxt_args (fn ctxt => |
|
150 |
Method.METHOD (fn facts => |
|
| 15032 | 151 |
fuf_tac (local_clasimpset_of ctxt) 1)) *} |
| 14299 | 152 |
"free ultrafilter tactic" |
153 |
||
154 |
method_setup ultra = {*
|
|
155 |
Method.ctxt_args (fn ctxt => |
|
156 |
Method.METHOD (fn facts => |
|
| 15032 | 157 |
ultra_tac (local_clasimpset_of ctxt) 1)) *} |
| 14299 | 158 |
"ultrafilter tactic" |
159 |
||
160 |
||
| 14329 | 161 |
text{*One further property of our free ultrafilter*}
|
162 |
lemma FreeUltrafilterNat_Un: |
|
| 14705 | 163 |
"X Un Y \<in> FreeUltrafilterNat |
164 |
==> X \<in> FreeUltrafilterNat | Y \<in> FreeUltrafilterNat" |
|
165 |
by (auto, ultra) |
|
| 14299 | 166 |
|
167 |
||
| 17298 | 168 |
subsection{*Properties of @{term starrel}*}
|
| 14329 | 169 |
|
| 17298 | 170 |
text{*Proving that @{term starrel} is an equivalence relation*}
|
| 14299 | 171 |
|
| 17298 | 172 |
lemma starrel_iff: "((X,Y) \<in> starrel) = ({n. X n = Y n} \<in> FreeUltrafilterNat)"
|
173 |
by (simp add: starrel_def) |
|
| 14299 | 174 |
|
| 17298 | 175 |
lemma starrel_refl: "(x,x) \<in> starrel" |
176 |
by (simp add: starrel_def) |
|
| 14299 | 177 |
|
| 17298 | 178 |
lemma starrel_sym [rule_format (no_asm)]: "(x,y) \<in> starrel --> (y,x) \<in> starrel" |
179 |
by (simp add: starrel_def eq_commute) |
|
| 14299 | 180 |
|
| 17298 | 181 |
lemma starrel_trans: |
182 |
"[|(x,y) \<in> starrel; (y,z) \<in> starrel|] ==> (x,z) \<in> starrel" |
|
183 |
by (simp add: starrel_def, ultra) |
|
| 14299 | 184 |
|
| 17298 | 185 |
lemma equiv_starrel: "equiv UNIV starrel" |
186 |
by (rule StarType.equiv_starrel) |
|
| 14299 | 187 |
|
| 17298 | 188 |
(* (starrel `` {x} = starrel `` {y}) = ((x,y) \<in> starrel) *)
|
189 |
lemmas equiv_starrel_iff = |
|
190 |
eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I, simp] |
|
| 14299 | 191 |
|
| 17298 | 192 |
lemma starrel_in_hypreal [simp]: "starrel``{x}:star"
|
193 |
by (simp add: star_def starrel_def quotient_def, blast) |
|
| 14299 | 194 |
|
| 17298 | 195 |
declare Abs_star_inject [simp] Abs_star_inverse [simp] |
196 |
declare equiv_starrel [THEN eq_equiv_class_iff, simp] |
|
197 |
declare starrel_iff [iff] |
|
| 14299 | 198 |
|
| 17298 | 199 |
lemmas eq_starrelD = eq_equiv_class [OF _ equiv_starrel] |
| 14299 | 200 |
|
| 17298 | 201 |
lemma lemma_starrel_refl [simp]: "x \<in> starrel `` {x}"
|
202 |
by (simp add: starrel_def) |
|
| 14299 | 203 |
|
| 17298 | 204 |
lemma hypreal_empty_not_mem [simp]: "{} \<notin> star"
|
205 |
apply (simp add: star_def) |
|
| 14299 | 206 |
apply (auto elim!: quotientE equalityCE) |
207 |
done |
|
208 |
||
| 17298 | 209 |
lemma Rep_hypreal_nonempty [simp]: "Rep_star x \<noteq> {}"
|
210 |
by (insert Rep_star [of x], auto) |
|
| 14299 | 211 |
|
| 14329 | 212 |
subsection{*@{term hypreal_of_real}:
|
213 |
the Injection from @{typ real} to @{typ hypreal}*}
|
|
| 14299 | 214 |
|
215 |
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
|
216 |
by (rule inj_onI, simp) |
| 14299 | 217 |
|
| 17298 | 218 |
lemma eq_Abs_star: |
219 |
"(!!x. z = Abs_star(starrel``{x}) ==> P) ==> P"
|
|
220 |
by (fold star_n_def, auto intro: star_cases) |
|
| 14299 | 221 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
222 |
lemma Rep_star_star_n_iff [simp]: |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
223 |
"(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
|
224 |
by (simp add: star_n_def) |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
225 |
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
226 |
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
|
227 |
by simp |
| 14329 | 228 |
|
229 |
subsection{*Hyperreal Addition*}
|
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
230 |
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
231 |
lemma star_n_add: |
| 17298 | 232 |
"star_n X + star_n Y = star_n (%n. X n + Y n)" |
233 |
by (simp add: star_add_def Ifun2_of_def star_of_def Ifun_star_n) |
|
| 14329 | 234 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
235 |
subsection{*Additive inverse on @{typ hypreal}*}
|
| 14329 | 236 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
237 |
lemma star_n_minus: |
| 17298 | 238 |
"- star_n X = star_n (%n. -(X n))" |
239 |
by (simp add: star_minus_def Ifun_of_def star_of_def Ifun_star_n) |
|
| 14299 | 240 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
241 |
lemma star_n_diff: |
| 17298 | 242 |
"star_n X - star_n Y = star_n (%n. X n - Y n)" |
243 |
by (simp add: star_diff_def Ifun2_of_def star_of_def Ifun_star_n) |
|
| 14299 | 244 |
|
| 14329 | 245 |
|
246 |
subsection{*Hyperreal Multiplication*}
|
|
| 14299 | 247 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
248 |
lemma star_n_mult: |
| 17298 | 249 |
"star_n X * star_n Y = star_n (%n. X n * Y n)" |
250 |
by (simp add: star_mult_def Ifun2_of_def star_of_def Ifun_star_n) |
|
| 14299 | 251 |
|
252 |
||
| 14329 | 253 |
subsection{*Multiplicative Inverse on @{typ hypreal} *}
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
254 |
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
255 |
lemma star_n_inverse: |
| 17298 | 256 |
"inverse (star_n X) = star_n (%n. if X n = (0::real) then 0 else inverse(X n))" |
257 |
apply (simp add: star_inverse_def Ifun_of_def star_of_def Ifun_star_n) |
|
258 |
apply (rule_tac f=star_n in arg_cong) |
|
259 |
apply (rule ext) |
|
260 |
apply simp |
|
261 |
done |
|
| 14299 | 262 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
263 |
lemma star_n_inverse2: |
| 17301 | 264 |
"inverse (star_n X) = star_n (%n. inverse(X n))" |
265 |
by (simp add: star_inverse_def Ifun_of_def star_of_def Ifun_star_n) |
|
266 |
||
| 14329 | 267 |
|
268 |
subsection{*Properties of The @{text "\<le>"} Relation*}
|
|
| 14299 | 269 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
270 |
lemma star_n_le: |
| 17298 | 271 |
"star_n X \<le> star_n Y = |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
272 |
({n. X n \<le> Y n} \<in> FreeUltrafilterNat)"
|
| 17298 | 273 |
by (simp add: star_le_def Ipred2_of_def star_of_def Ifun_star_n unstar_star_n) |
| 14299 | 274 |
|
| 14370 | 275 |
lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \<noteq> y" |
| 15539 | 276 |
by (auto) |
| 14329 | 277 |
|
278 |
subsection{*The Hyperreals Form an Ordered Field*}
|
|
279 |
||
| 14331 | 280 |
lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)" |
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
281 |
by auto |
| 14331 | 282 |
|
| 14329 | 283 |
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
|
284 |
by auto |
| 14329 | 285 |
|
286 |
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
|
287 |
by auto |
| 14329 | 288 |
|
289 |
||
290 |
subsection{*Misc Others*}
|
|
| 14299 | 291 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
292 |
lemma star_n_less: |
| 17298 | 293 |
"star_n X < star_n Y = ({n. X n < Y n} \<in> FreeUltrafilterNat)"
|
294 |
by (simp add: star_less_def Ipred2_of_def star_of_def Ifun_star_n unstar_star_n) |
|
| 14370 | 295 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
296 |
lemma star_n_zero_num: "0 = star_n (%n. 0)" |
| 17298 | 297 |
by (simp add: star_zero_def star_of_def) |
| 14299 | 298 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
299 |
lemma star_n_one_num: "1 = star_n (%n. 1)" |
| 17298 | 300 |
by (simp add: star_one_def star_of_def) |
| 14299 | 301 |
|
| 14301 | 302 |
lemma hypreal_omega_gt_zero [simp]: "0 < omega" |
| 17298 | 303 |
apply (simp only: omega_def star_zero_def star_less_def star_of_def) |
304 |
apply (simp add: Ipred2_of_def star_of_def Ifun_star_n unstar_star_n) |
|
| 14329 | 305 |
done |
306 |
||
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
307 |
lemma star_n_abs: |
| 17298 | 308 |
"abs (star_n X) = star_n (%n. abs (X n))" |
309 |
by (simp add: star_abs_def Ifun_of_def star_of_def Ifun_star_n) |
|
| 14370 | 310 |
|
311 |
subsection{*Existence of Infinite Hyperreal Number*}
|
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
312 |
|
| 14370 | 313 |
text{*Existence of infinite number not corresponding to any real number.
|
314 |
Use assumption that member @{term FreeUltrafilterNat} is not finite.*}
|
|
315 |
||
316 |
||
317 |
text{*A few lemmas first*}
|
|
318 |
||
319 |
lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} |
|
|
320 |
(\<exists>y. {n::nat. x = real n} = {y})"
|
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
321 |
by force |
| 14370 | 322 |
|
323 |
lemma lemma_finite_omega_set: "finite {n::nat. x = real n}"
|
|
324 |
by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto) |
|
325 |
||
326 |
lemma not_ex_hypreal_of_real_eq_omega: |
|
327 |
"~ (\<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
|
328 |
apply (simp add: omega_def) |
| 17298 | 329 |
apply (simp add: star_of_def star_n_eq_iff) |
| 14370 | 330 |
apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] |
331 |
lemma_finite_omega_set [THEN FreeUltrafilterNat_finite]) |
|
332 |
done |
|
333 |
||
334 |
lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \<noteq> omega" |
|
| 14705 | 335 |
by (insert not_ex_hypreal_of_real_eq_omega, auto) |
| 14370 | 336 |
|
337 |
text{*Existence of infinitesimal number also not corresponding to any
|
|
338 |
real number*} |
|
339 |
||
340 |
lemma lemma_epsilon_empty_singleton_disj: |
|
341 |
"{n::nat. x = inverse(real(Suc n))} = {} |
|
|
342 |
(\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})"
|
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
343 |
by auto |
| 14370 | 344 |
|
345 |
lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}"
|
|
346 |
by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto) |
|
347 |
||
| 14705 | 348 |
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
|
349 |
by (auto simp add: epsilon_def star_of_def star_n_eq_iff |
| 14705 | 350 |
lemma_finite_epsilon_set [THEN FreeUltrafilterNat_finite]) |
| 14370 | 351 |
|
352 |
lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> epsilon" |
|
| 14705 | 353 |
by (insert not_ex_hypreal_of_real_eq_epsilon, auto) |
| 14370 | 354 |
|
355 |
lemma hypreal_epsilon_not_zero: "epsilon \<noteq> 0" |
|
| 17298 | 356 |
by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff |
357 |
del: star_of_zero) |
|
| 14370 | 358 |
|
359 |
lemma hypreal_epsilon_inverse_omega: "epsilon = inverse(omega)" |
|
| 17298 | 360 |
apply (simp add: epsilon_def omega_def star_inverse_def) |
361 |
apply (simp add: Ifun_of_def star_of_def Ifun_star_n) |
|
362 |
done |
|
| 14370 | 363 |
|
364 |
||
| 14299 | 365 |
ML |
366 |
{*
|
|
367 |
val omega_def = thm "omega_def"; |
|
368 |
val epsilon_def = thm "epsilon_def"; |
|
369 |
||
370 |
val finite_exhausts = thm "finite_exhausts"; |
|
371 |
val finite_not_covers = thm "finite_not_covers"; |
|
372 |
val not_finite_nat = thm "not_finite_nat"; |
|
373 |
val FreeUltrafilterNat_Ex = thm "FreeUltrafilterNat_Ex"; |
|
374 |
val FreeUltrafilterNat_mem = thm "FreeUltrafilterNat_mem"; |
|
375 |
val FreeUltrafilterNat_finite = thm "FreeUltrafilterNat_finite"; |
|
376 |
val FreeUltrafilterNat_not_finite = thm "FreeUltrafilterNat_not_finite"; |
|
377 |
val FreeUltrafilterNat_empty = thm "FreeUltrafilterNat_empty"; |
|
378 |
val FreeUltrafilterNat_Int = thm "FreeUltrafilterNat_Int"; |
|
379 |
val FreeUltrafilterNat_subset = thm "FreeUltrafilterNat_subset"; |
|
380 |
val FreeUltrafilterNat_Compl = thm "FreeUltrafilterNat_Compl"; |
|
381 |
val FreeUltrafilterNat_Compl_mem = thm "FreeUltrafilterNat_Compl_mem"; |
|
382 |
val FreeUltrafilterNat_Compl_iff1 = thm "FreeUltrafilterNat_Compl_iff1"; |
|
383 |
val FreeUltrafilterNat_Compl_iff2 = thm "FreeUltrafilterNat_Compl_iff2"; |
|
384 |
val FreeUltrafilterNat_UNIV = thm "FreeUltrafilterNat_UNIV"; |
|
385 |
val FreeUltrafilterNat_Nat_set_refl = thm "FreeUltrafilterNat_Nat_set_refl"; |
|
386 |
val FreeUltrafilterNat_P = thm "FreeUltrafilterNat_P"; |
|
387 |
val FreeUltrafilterNat_Ex_P = thm "FreeUltrafilterNat_Ex_P"; |
|
388 |
val FreeUltrafilterNat_all = thm "FreeUltrafilterNat_all"; |
|
389 |
val FreeUltrafilterNat_Un = thm "FreeUltrafilterNat_Un"; |
|
| 17298 | 390 |
val starrel_iff = thm "starrel_iff"; |
391 |
val starrel_in_hypreal = thm "starrel_in_hypreal"; |
|
392 |
val Abs_star_inverse = thm "Abs_star_inverse"; |
|
393 |
val lemma_starrel_refl = thm "lemma_starrel_refl"; |
|
| 14299 | 394 |
val hypreal_empty_not_mem = thm "hypreal_empty_not_mem"; |
395 |
val Rep_hypreal_nonempty = thm "Rep_hypreal_nonempty"; |
|
396 |
val inj_hypreal_of_real = thm "inj_hypreal_of_real"; |
|
| 17298 | 397 |
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
|
398 |
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
|
399 |
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
|
400 |
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
|
401 |
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
|
402 |
val star_n_inverse = thm "star_n_inverse"; |
| 14299 | 403 |
val hypreal_mult_left_cancel = thm "hypreal_mult_left_cancel"; |
404 |
val hypreal_mult_right_cancel = thm "hypreal_mult_right_cancel"; |
|
405 |
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
|
406 |
val star_n_less = thm "star_n_less"; |
| 14299 | 407 |
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
|
408 |
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
|
409 |
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
|
410 |
val star_n_one_num = thm "star_n_one_num"; |
| 14299 | 411 |
val hypreal_omega_gt_zero = thm "hypreal_omega_gt_zero"; |
| 14370 | 412 |
|
413 |
val lemma_omega_empty_singleton_disj = thm"lemma_omega_empty_singleton_disj"; |
|
414 |
val lemma_finite_omega_set = thm"lemma_finite_omega_set"; |
|
415 |
val not_ex_hypreal_of_real_eq_omega = thm"not_ex_hypreal_of_real_eq_omega"; |
|
416 |
val hypreal_of_real_not_eq_omega = thm"hypreal_of_real_not_eq_omega"; |
|
417 |
val not_ex_hypreal_of_real_eq_epsilon = thm"not_ex_hypreal_of_real_eq_epsilon"; |
|
418 |
val hypreal_of_real_not_eq_epsilon = thm"hypreal_of_real_not_eq_epsilon"; |
|
419 |
val hypreal_epsilon_not_zero = thm"hypreal_epsilon_not_zero"; |
|
420 |
val hypreal_epsilon_inverse_omega = thm"hypreal_epsilon_inverse_omega"; |
|
| 14299 | 421 |
*} |
422 |
||
| 10751 | 423 |
end |