| author | wenzelm |
| Wed, 13 Dec 2006 15:47:34 +0100 | |
| changeset 21824 | 153fad1e7318 |
| parent 21810 | b2d23672b003 |
| child 21855 | 74909ecaf20a |
| permissions | -rw-r--r-- |
| 20245 | 1 |
(* Title : HOL/Hyperreal/HyperDef.thy |
| 10751 | 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 |
|
| 19380 | 17 |
abbreviation |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
18 |
hypreal_of_real :: "real => real star" where |
| 19380 | 19 |
"hypreal_of_real == star_of" |
| 10751 | 20 |
|
| 19765 | 21 |
definition |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
22 |
omega :: hypreal where |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
23 |
-- {*an infinite number @{text "= [<1,2,3,...>]"} *}
|
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
24 |
"omega = star_n (\<lambda>n. real (Suc n))" |
| 13487 | 25 |
|
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
26 |
definition |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
27 |
epsilon :: hypreal where |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
28 |
-- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *}
|
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
29 |
"epsilon = star_n (\<lambda>n. inverse (real (Suc n)))" |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
30 |
|
| 21210 | 31 |
notation (xsymbols) |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
32 |
omega ("\<omega>") and
|
| 19765 | 33 |
epsilon ("\<epsilon>")
|
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
34 |
|
| 21210 | 35 |
notation (HTML output) |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
36 |
omega ("\<omega>") and
|
| 19765 | 37 |
epsilon ("\<epsilon>")
|
| 14565 | 38 |
|
| 14329 | 39 |
|
| 20555 | 40 |
subsection {* Real vector class instances *}
|
41 |
||
42 |
instance star :: (scaleR) scaleR .. |
|
43 |
||
44 |
defs (overloaded) |
|
45 |
star_scaleR_def [transfer_unfold]: "scaleR r \<equiv> *f* (scaleR r)" |
|
46 |
||
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
47 |
lemma Standard_scaleR [simp]: "x \<in> Standard \<Longrightarrow> scaleR r x \<in> Standard" |
| 20726 | 48 |
by (simp add: star_scaleR_def) |
49 |
||
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
50 |
lemma star_of_scaleR [simp]: "star_of (scaleR r x) = scaleR r (star_of x)" |
| 20726 | 51 |
by transfer (rule refl) |
52 |
||
| 20555 | 53 |
instance star :: (real_vector) real_vector |
54 |
proof |
|
55 |
fix a b :: real |
|
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
56 |
show "\<And>x y::'a star. scaleR a (x + y) = scaleR a x + scaleR a y" |
| 20555 | 57 |
by transfer (rule scaleR_right_distrib) |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
58 |
show "\<And>x::'a star. scaleR (a + b) x = scaleR a x + scaleR b x" |
| 20555 | 59 |
by transfer (rule scaleR_left_distrib) |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
60 |
show "\<And>x::'a star. scaleR a (scaleR b x) = scaleR (a * b) x" |
| 20765 | 61 |
by transfer (rule scaleR_scaleR) |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
62 |
show "\<And>x::'a star. scaleR 1 x = x" |
| 20555 | 63 |
by transfer (rule scaleR_one) |
64 |
qed |
|
65 |
||
66 |
instance star :: (real_algebra) real_algebra |
|
67 |
proof |
|
68 |
fix a :: real |
|
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
69 |
show "\<And>x y::'a star. scaleR a x * y = scaleR a (x * y)" |
| 20555 | 70 |
by transfer (rule mult_scaleR_left) |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
71 |
show "\<And>x y::'a star. x * scaleR a y = scaleR a (x * y)" |
| 20555 | 72 |
by transfer (rule mult_scaleR_right) |
73 |
qed |
|
74 |
||
75 |
instance star :: (real_algebra_1) real_algebra_1 .. |
|
76 |
||
|
20584
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
huffman
parents:
20555
diff
changeset
|
77 |
instance star :: (real_div_algebra) real_div_algebra .. |
|
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
huffman
parents:
20555
diff
changeset
|
78 |
|
|
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
huffman
parents:
20555
diff
changeset
|
79 |
instance star :: (real_field) real_field .. |
|
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
huffman
parents:
20555
diff
changeset
|
80 |
|
| 20726 | 81 |
lemma star_of_real_def [transfer_unfold]: "of_real r = star_of (of_real r)" |
82 |
by (unfold of_real_def, transfer, rule refl) |
|
83 |
||
84 |
lemma Standard_of_real [simp]: "of_real r \<in> Standard" |
|
85 |
by (simp add: star_of_real_def) |
|
| 20555 | 86 |
|
87 |
lemma star_of_of_real [simp]: "star_of (of_real r) = of_real r" |
|
88 |
by transfer (rule refl) |
|
89 |
||
| 20729 | 90 |
lemma of_real_eq_star_of [simp]: "of_real = star_of" |
91 |
proof |
|
92 |
fix r :: real |
|
93 |
show "of_real r = star_of r" |
|
94 |
by transfer simp |
|
95 |
qed |
|
96 |
||
97 |
lemma Reals_eq_Standard: "(Reals :: hypreal set) = Standard" |
|
| 20765 | 98 |
by (simp add: Reals_def Standard_def) |
| 20729 | 99 |
|
| 20555 | 100 |
|
| 14329 | 101 |
subsection{*Existence of Free Ultrafilter over the Naturals*}
|
102 |
||
103 |
text{*Also, proof of various properties of @{term FreeUltrafilterNat}:
|
|
104 |
an arbitrary free ultrafilter*} |
|
| 14299 | 105 |
|
| 17290 | 106 |
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
|
107 |
by (rule nat_infinite [THEN freeultrafilter_Ex]) |
| 14299 | 108 |
|
| 17298 | 109 |
lemma FreeUltrafilterNat_mem: "freeultrafilter FreeUltrafilterNat" |
| 14299 | 110 |
apply (unfold FreeUltrafilterNat_def) |
| 17290 | 111 |
apply (rule someI_ex) |
112 |
apply (rule FreeUltrafilterNat_Ex) |
|
| 14299 | 113 |
done |
114 |
||
| 17290 | 115 |
lemma UltrafilterNat_mem: "ultrafilter FreeUltrafilterNat" |
116 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.ultrafilter]) |
|
117 |
||
118 |
lemma FilterNat_mem: "filter FreeUltrafilterNat" |
|
119 |
by (rule FreeUltrafilterNat_mem [THEN freeultrafilter.filter]) |
|
120 |
||
| 14299 | 121 |
lemma FreeUltrafilterNat_finite: "finite x ==> x \<notin> FreeUltrafilterNat" |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
122 |
by (rule FreeUltrafilterNat.finite) |
| 14299 | 123 |
|
| 14705 | 124 |
lemma FreeUltrafilterNat_not_finite: "x \<in> FreeUltrafilterNat ==> ~ finite x" |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
125 |
by (rule FreeUltrafilterNat.infinite) |
| 14299 | 126 |
|
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
127 |
lemma FreeUltrafilterNat_empty: "{} \<notin> FreeUltrafilterNat"
|
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
128 |
by (rule FreeUltrafilterNat.empty) |
| 14299 | 129 |
|
| 14329 | 130 |
lemma FreeUltrafilterNat_Int: |
| 14705 | 131 |
"[| X \<in> FreeUltrafilterNat; Y \<in> FreeUltrafilterNat |] |
| 14299 | 132 |
==> X Int Y \<in> FreeUltrafilterNat" |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
133 |
by (rule FreeUltrafilterNat.Int) |
| 14299 | 134 |
|
| 14329 | 135 |
lemma FreeUltrafilterNat_subset: |
| 14705 | 136 |
"[| X \<in> FreeUltrafilterNat; X \<subseteq> Y |] |
| 14299 | 137 |
==> Y \<in> FreeUltrafilterNat" |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
138 |
by (rule FreeUltrafilterNat.subset) |
| 14299 | 139 |
|
| 14329 | 140 |
lemma FreeUltrafilterNat_Compl: |
| 14705 | 141 |
"X \<in> FreeUltrafilterNat ==> -X \<notin> FreeUltrafilterNat" |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
142 |
by (rule FreeUltrafilterNat.memD) |
| 14299 | 143 |
|
| 14329 | 144 |
lemma FreeUltrafilterNat_Compl_mem: |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
145 |
"X \<notin> FreeUltrafilterNat ==> -X \<in> FreeUltrafilterNat" |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
146 |
by (rule FreeUltrafilterNat.not_memD) |
| 14299 | 147 |
|
| 14329 | 148 |
lemma FreeUltrafilterNat_Compl_iff1: |
| 14705 | 149 |
"(X \<notin> FreeUltrafilterNat) = (-X \<in> FreeUltrafilterNat)" |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
150 |
by (rule FreeUltrafilterNat.not_mem_iff) |
| 14299 | 151 |
|
| 14329 | 152 |
lemma FreeUltrafilterNat_Compl_iff2: |
| 14705 | 153 |
"(X \<in> FreeUltrafilterNat) = (-X \<notin> FreeUltrafilterNat)" |
| 14301 | 154 |
by (auto simp add: FreeUltrafilterNat_Compl_iff1 [symmetric]) |
| 14299 | 155 |
|
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
156 |
lemma cofinite_mem_FreeUltrafilterNat: "finite (-X) ==> X \<in> FreeUltrafilterNat" |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
157 |
apply (drule FreeUltrafilterNat.finite) |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
158 |
apply (simp add: FreeUltrafilterNat.not_mem_iff) |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
159 |
done |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
160 |
|
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
161 |
lemma FreeUltrafilterNat_UNIV: "UNIV \<in> FreeUltrafilterNat" |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
162 |
by (rule FreeUltrafilterNat.UNIV) |
| 14299 | 163 |
|
| 14329 | 164 |
lemma FreeUltrafilterNat_Nat_set_refl [intro]: |
165 |
"{n. P(n) = P(n)} \<in> FreeUltrafilterNat"
|
|
| 14301 | 166 |
by simp |
| 14299 | 167 |
|
168 |
lemma FreeUltrafilterNat_P: "{n::nat. P} \<in> FreeUltrafilterNat ==> P"
|
|
| 14301 | 169 |
by (rule ccontr, simp) |
| 14299 | 170 |
|
171 |
lemma FreeUltrafilterNat_Ex_P: "{n. P(n)} \<in> FreeUltrafilterNat ==> \<exists>n. P(n)"
|
|
| 14301 | 172 |
by (rule ccontr, simp) |
| 14299 | 173 |
|
174 |
lemma FreeUltrafilterNat_all: "\<forall>n. P(n) ==> {n. P(n)} \<in> FreeUltrafilterNat"
|
|
| 15539 | 175 |
by (auto) |
| 14299 | 176 |
|
| 14329 | 177 |
|
178 |
text{*Define and use Ultrafilter tactics*}
|
|
| 14299 | 179 |
use "fuf.ML" |
180 |
||
181 |
method_setup fuf = {*
|
|
182 |
Method.ctxt_args (fn ctxt => |
|
| 21588 | 183 |
Method.SIMPLE_METHOD' (fuf_tac (local_clasimpset_of ctxt))) *} |
| 14299 | 184 |
"free ultrafilter tactic" |
185 |
||
186 |
method_setup ultra = {*
|
|
187 |
Method.ctxt_args (fn ctxt => |
|
| 21588 | 188 |
Method.SIMPLE_METHOD' (ultra_tac (local_clasimpset_of ctxt))) *} |
| 14299 | 189 |
"ultrafilter tactic" |
190 |
||
191 |
||
| 14329 | 192 |
text{*One further property of our free ultrafilter*}
|
193 |
lemma FreeUltrafilterNat_Un: |
|
| 14705 | 194 |
"X Un Y \<in> FreeUltrafilterNat |
195 |
==> X \<in> FreeUltrafilterNat | Y \<in> FreeUltrafilterNat" |
|
196 |
by (auto, ultra) |
|
| 14299 | 197 |
|
198 |
||
| 17298 | 199 |
subsection{*Properties of @{term starrel}*}
|
| 14329 | 200 |
|
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
201 |
lemma lemma_starrel_refl [simp]: "x \<in> starrel `` {x}"
|
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
202 |
by (simp add: starrel_def) |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
203 |
|
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
204 |
lemma starrel_in_hypreal [simp]: "starrel``{x}:star"
|
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
205 |
by (simp add: star_def starrel_def quotient_def, blast) |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
206 |
|
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
207 |
declare Abs_star_inject [simp] Abs_star_inverse [simp] |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
208 |
declare equiv_starrel [THEN eq_equiv_class_iff, simp] |
| 14299 | 209 |
|
| 14329 | 210 |
subsection{*@{term hypreal_of_real}:
|
211 |
the Injection from @{typ real} to @{typ hypreal}*}
|
|
| 14299 | 212 |
|
|
21810
b2d23672b003
generalized some lemmas; removed redundant lemmas; cleaned up some proofs
huffman
parents:
21787
diff
changeset
|
213 |
lemma inj_star_of: "inj star_of" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
214 |
by (rule inj_onI, simp) |
| 14299 | 215 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20245
diff
changeset
|
216 |
lemma mem_Rep_star_iff: "(X \<in> Rep_star x) = (x = star_n X)" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20245
diff
changeset
|
217 |
by (cases x, simp add: star_n_def) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20245
diff
changeset
|
218 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
219 |
lemma Rep_star_star_n_iff [simp]: |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
220 |
"(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
|
221 |
by (simp add: star_n_def) |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
222 |
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
223 |
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
|
224 |
by simp |
| 14329 | 225 |
|
|
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
226 |
subsection{* Properties of @{term star_n} *}
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
227 |
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
228 |
lemma star_n_add: |
| 17298 | 229 |
"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
|
230 |
by (simp only: star_add_def starfun2_star_n) |
| 14329 | 231 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
232 |
lemma star_n_minus: |
| 17298 | 233 |
"- 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
|
234 |
by (simp only: star_minus_def starfun_star_n) |
| 14299 | 235 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
236 |
lemma star_n_diff: |
| 17298 | 237 |
"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
|
238 |
by (simp only: star_diff_def starfun2_star_n) |
| 14299 | 239 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
240 |
lemma star_n_mult: |
| 17298 | 241 |
"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
|
242 |
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
|
243 |
|
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
244 |
lemma star_n_inverse: |
| 17301 | 245 |
"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
|
246 |
by (simp only: star_inverse_def starfun_star_n) |
| 14299 | 247 |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
248 |
lemma star_n_le: |
| 17298 | 249 |
"star_n X \<le> star_n Y = |
|
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
250 |
({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
|
251 |
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
|
252 |
|
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
253 |
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
|
254 |
"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
|
255 |
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
|
256 |
|
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
257 |
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
|
258 |
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
|
259 |
|
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
260 |
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
|
261 |
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
|
262 |
|
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
263 |
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
|
264 |
"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
|
265 |
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
|
266 |
|
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
267 |
subsection{*Misc Others*}
|
| 14299 | 268 |
|
| 14370 | 269 |
lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \<noteq> y" |
| 15539 | 270 |
by (auto) |
| 14329 | 271 |
|
| 14331 | 272 |
lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)" |
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
273 |
by auto |
| 14331 | 274 |
|
| 14329 | 275 |
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
|
276 |
by auto |
| 14329 | 277 |
|
278 |
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
|
279 |
by auto |
| 14329 | 280 |
|
| 14301 | 281 |
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
|
282 |
by (simp add: omega_def star_n_zero_num star_n_less) |
| 14370 | 283 |
|
284 |
subsection{*Existence of Infinite Hyperreal Number*}
|
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
285 |
|
| 14370 | 286 |
text{*Existence of infinite number not corresponding to any real number.
|
287 |
Use assumption that member @{term FreeUltrafilterNat} is not finite.*}
|
|
288 |
||
289 |
||
290 |
text{*A few lemmas first*}
|
|
291 |
||
292 |
lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} |
|
|
293 |
(\<exists>y. {n::nat. x = real n} = {y})"
|
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
294 |
by force |
| 14370 | 295 |
|
296 |
lemma lemma_finite_omega_set: "finite {n::nat. x = real n}"
|
|
297 |
by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto) |
|
298 |
||
299 |
lemma not_ex_hypreal_of_real_eq_omega: |
|
300 |
"~ (\<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
|
301 |
apply (simp add: omega_def) |
| 17298 | 302 |
apply (simp add: star_of_def star_n_eq_iff) |
| 14370 | 303 |
apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] |
304 |
lemma_finite_omega_set [THEN FreeUltrafilterNat_finite]) |
|
305 |
done |
|
306 |
||
307 |
lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \<noteq> omega" |
|
| 14705 | 308 |
by (insert not_ex_hypreal_of_real_eq_omega, auto) |
| 14370 | 309 |
|
310 |
text{*Existence of infinitesimal number also not corresponding to any
|
|
311 |
real number*} |
|
312 |
||
313 |
lemma lemma_epsilon_empty_singleton_disj: |
|
314 |
"{n::nat. x = inverse(real(Suc n))} = {} |
|
|
315 |
(\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})"
|
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
316 |
by auto |
| 14370 | 317 |
|
318 |
lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}"
|
|
319 |
by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto) |
|
320 |
||
| 14705 | 321 |
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
|
322 |
by (auto simp add: epsilon_def star_of_def star_n_eq_iff |
| 14705 | 323 |
lemma_finite_epsilon_set [THEN FreeUltrafilterNat_finite]) |
| 14370 | 324 |
|
325 |
lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> epsilon" |
|
| 14705 | 326 |
by (insert not_ex_hypreal_of_real_eq_epsilon, auto) |
| 14370 | 327 |
|
328 |
lemma hypreal_epsilon_not_zero: "epsilon \<noteq> 0" |
|
| 17298 | 329 |
by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff |
330 |
del: star_of_zero) |
|
| 14370 | 331 |
|
332 |
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
|
333 |
by (simp add: epsilon_def omega_def star_n_inverse) |
| 14370 | 334 |
|
| 20753 | 335 |
lemma hypreal_epsilon_gt_zero: "0 < epsilon" |
336 |
by (simp add: hypreal_epsilon_inverse_omega) |
|
337 |
||
| 10751 | 338 |
end |