author | huffman |
Thu, 14 Dec 2006 22:18:08 +0100 | |
changeset 21856 | f44628fb2033 |
parent 21855 | 74909ecaf20a |
child 21865 | 55cc354fd2d9 |
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" |
15131 | 12 |
begin |
10751 | 13 |
|
17298 | 14 |
types hypreal = "real star" |
14299 | 15 |
|
19380 | 16 |
abbreviation |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
17 |
hypreal_of_real :: "real => real star" where |
19380 | 18 |
"hypreal_of_real == star_of" |
10751 | 19 |
|
19765 | 20 |
definition |
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
21 |
omega :: hypreal where |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
22 |
-- {*an infinite number @{text "= [<1,2,3,...>]"} *} |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
23 |
"omega = star_n (\<lambda>n. real (Suc n))" |
13487 | 24 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
25 |
definition |
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
26 |
epsilon :: hypreal where |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
27 |
-- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *} |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
28 |
"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
|
29 |
|
21210 | 30 |
notation (xsymbols) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
31 |
omega ("\<omega>") and |
19765 | 32 |
epsilon ("\<epsilon>") |
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
33 |
|
21210 | 34 |
notation (HTML output) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
35 |
omega ("\<omega>") and |
19765 | 36 |
epsilon ("\<epsilon>") |
14565 | 37 |
|
14329 | 38 |
|
20555 | 39 |
subsection {* Real vector class instances *} |
40 |
||
41 |
instance star :: (scaleR) scaleR .. |
|
42 |
||
43 |
defs (overloaded) |
|
44 |
star_scaleR_def [transfer_unfold]: "scaleR r \<equiv> *f* (scaleR r)" |
|
45 |
||
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
46 |
lemma Standard_scaleR [simp]: "x \<in> Standard \<Longrightarrow> scaleR r x \<in> Standard" |
20726 | 47 |
by (simp add: star_scaleR_def) |
48 |
||
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
49 |
lemma star_of_scaleR [simp]: "star_of (scaleR r x) = scaleR r (star_of x)" |
20726 | 50 |
by transfer (rule refl) |
51 |
||
20555 | 52 |
instance star :: (real_vector) real_vector |
53 |
proof |
|
54 |
fix a b :: real |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
55 |
show "\<And>x y::'a star. scaleR a (x + y) = scaleR a x + scaleR a y" |
20555 | 56 |
by transfer (rule scaleR_right_distrib) |
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
57 |
show "\<And>x::'a star. scaleR (a + b) x = scaleR a x + scaleR b x" |
20555 | 58 |
by transfer (rule scaleR_left_distrib) |
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
59 |
show "\<And>x::'a star. scaleR a (scaleR b x) = scaleR (a * b) x" |
20765 | 60 |
by transfer (rule scaleR_scaleR) |
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
61 |
show "\<And>x::'a star. scaleR 1 x = x" |
20555 | 62 |
by transfer (rule scaleR_one) |
63 |
qed |
|
64 |
||
65 |
instance star :: (real_algebra) real_algebra |
|
66 |
proof |
|
67 |
fix a :: real |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
68 |
show "\<And>x y::'a star. scaleR a x * y = scaleR a (x * y)" |
20555 | 69 |
by transfer (rule mult_scaleR_left) |
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
70 |
show "\<And>x y::'a star. x * scaleR a y = scaleR a (x * y)" |
20555 | 71 |
by transfer (rule mult_scaleR_right) |
72 |
qed |
|
73 |
||
74 |
instance star :: (real_algebra_1) real_algebra_1 .. |
|
75 |
||
20584
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
huffman
parents:
20555
diff
changeset
|
76 |
instance star :: (real_div_algebra) real_div_algebra .. |
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
huffman
parents:
20555
diff
changeset
|
77 |
|
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
huffman
parents:
20555
diff
changeset
|
78 |
instance star :: (real_field) real_field .. |
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
huffman
parents:
20555
diff
changeset
|
79 |
|
20726 | 80 |
lemma star_of_real_def [transfer_unfold]: "of_real r = star_of (of_real r)" |
81 |
by (unfold of_real_def, transfer, rule refl) |
|
82 |
||
83 |
lemma Standard_of_real [simp]: "of_real r \<in> Standard" |
|
84 |
by (simp add: star_of_real_def) |
|
20555 | 85 |
|
86 |
lemma star_of_of_real [simp]: "star_of (of_real r) = of_real r" |
|
87 |
by transfer (rule refl) |
|
88 |
||
20729 | 89 |
lemma of_real_eq_star_of [simp]: "of_real = star_of" |
90 |
proof |
|
91 |
fix r :: real |
|
92 |
show "of_real r = star_of r" |
|
93 |
by transfer simp |
|
94 |
qed |
|
95 |
||
96 |
lemma Reals_eq_Standard: "(Reals :: hypreal set) = Standard" |
|
20765 | 97 |
by (simp add: Reals_def Standard_def) |
20729 | 98 |
|
20555 | 99 |
|
17298 | 100 |
subsection{*Properties of @{term starrel}*} |
14329 | 101 |
|
21787
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
102 |
lemma lemma_starrel_refl [simp]: "x \<in> starrel `` {x}" |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
103 |
by (simp add: starrel_def) |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
104 |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
105 |
lemma starrel_in_hypreal [simp]: "starrel``{x}:star" |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
106 |
by (simp add: star_def starrel_def quotient_def, blast) |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
107 |
|
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
108 |
declare Abs_star_inject [simp] Abs_star_inverse [simp] |
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
huffman
parents:
21588
diff
changeset
|
109 |
declare equiv_starrel [THEN eq_equiv_class_iff, simp] |
14299 | 110 |
|
14329 | 111 |
subsection{*@{term hypreal_of_real}: |
112 |
the Injection from @{typ real} to @{typ hypreal}*} |
|
14299 | 113 |
|
21810
b2d23672b003
generalized some lemmas; removed redundant lemmas; cleaned up some proofs
huffman
parents:
21787
diff
changeset
|
114 |
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
|
115 |
by (rule inj_onI, simp) |
14299 | 116 |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20245
diff
changeset
|
117 |
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
|
118 |
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
|
119 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
120 |
lemma Rep_star_star_n_iff [simp]: |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
121 |
"(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
|
122 |
by (simp add: star_n_def) |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
123 |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
124 |
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
|
125 |
by simp |
14329 | 126 |
|
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
127 |
subsection{* Properties of @{term star_n} *} |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
128 |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
129 |
lemma star_n_add: |
17298 | 130 |
"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
|
131 |
by (simp only: star_add_def starfun2_star_n) |
14329 | 132 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
133 |
lemma star_n_minus: |
17298 | 134 |
"- 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
|
135 |
by (simp only: star_minus_def starfun_star_n) |
14299 | 136 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
137 |
lemma star_n_diff: |
17298 | 138 |
"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
|
139 |
by (simp only: star_diff_def starfun2_star_n) |
14299 | 140 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
141 |
lemma star_n_mult: |
17298 | 142 |
"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
|
143 |
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
|
144 |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
145 |
lemma star_n_inverse: |
17301 | 146 |
"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
|
147 |
by (simp only: star_inverse_def starfun_star_n) |
14299 | 148 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
149 |
lemma star_n_le: |
17298 | 150 |
"star_n X \<le> star_n Y = |
14365
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents:
14361
diff
changeset
|
151 |
({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
|
152 |
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
|
153 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
154 |
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
|
155 |
"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
|
156 |
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
|
157 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
158 |
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
|
159 |
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
|
160 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
161 |
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
|
162 |
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
|
163 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
164 |
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
|
165 |
"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
|
166 |
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
|
167 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17318
diff
changeset
|
168 |
subsection{*Misc Others*} |
14299 | 169 |
|
14370 | 170 |
lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \<noteq> y" |
15539 | 171 |
by (auto) |
14329 | 172 |
|
14331 | 173 |
lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)" |
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
174 |
by auto |
14331 | 175 |
|
14329 | 176 |
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
|
177 |
by auto |
14329 | 178 |
|
179 |
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
|
180 |
by auto |
14329 | 181 |
|
14301 | 182 |
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
|
183 |
by (simp add: omega_def star_n_zero_num star_n_less) |
14370 | 184 |
|
185 |
subsection{*Existence of Infinite Hyperreal Number*} |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17301
diff
changeset
|
186 |
|
14370 | 187 |
text{*Existence of infinite number not corresponding to any real number. |
188 |
Use assumption that member @{term FreeUltrafilterNat} is not finite.*} |
|
189 |
||
190 |
||
191 |
text{*A few lemmas first*} |
|
192 |
||
193 |
lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} | |
|
194 |
(\<exists>y. {n::nat. x = real n} = {y})" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
195 |
by force |
14370 | 196 |
|
197 |
lemma lemma_finite_omega_set: "finite {n::nat. x = real n}" |
|
198 |
by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto) |
|
199 |
||
200 |
lemma not_ex_hypreal_of_real_eq_omega: |
|
201 |
"~ (\<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
|
202 |
apply (simp add: omega_def) |
17298 | 203 |
apply (simp add: star_of_def star_n_eq_iff) |
14370 | 204 |
apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] |
21855
74909ecaf20a
remove ultra tactic and redundant FreeUltrafilterNat lemmas
huffman
parents:
21810
diff
changeset
|
205 |
lemma_finite_omega_set [THEN FreeUltrafilterNat.finite]) |
14370 | 206 |
done |
207 |
||
208 |
lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \<noteq> omega" |
|
14705 | 209 |
by (insert not_ex_hypreal_of_real_eq_omega, auto) |
14370 | 210 |
|
211 |
text{*Existence of infinitesimal number also not corresponding to any |
|
212 |
real number*} |
|
213 |
||
214 |
lemma lemma_epsilon_empty_singleton_disj: |
|
215 |
"{n::nat. x = inverse(real(Suc n))} = {} | |
|
216 |
(\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
217 |
by auto |
14370 | 218 |
|
219 |
lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}" |
|
220 |
by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto) |
|
221 |
||
14705 | 222 |
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
|
223 |
by (auto simp add: epsilon_def star_of_def star_n_eq_iff |
21855
74909ecaf20a
remove ultra tactic and redundant FreeUltrafilterNat lemmas
huffman
parents:
21810
diff
changeset
|
224 |
lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite]) |
14370 | 225 |
|
226 |
lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> epsilon" |
|
14705 | 227 |
by (insert not_ex_hypreal_of_real_eq_epsilon, auto) |
14370 | 228 |
|
229 |
lemma hypreal_epsilon_not_zero: "epsilon \<noteq> 0" |
|
17298 | 230 |
by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff |
231 |
del: star_of_zero) |
|
14370 | 232 |
|
233 |
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
|
234 |
by (simp add: epsilon_def omega_def star_n_inverse) |
14370 | 235 |
|
20753 | 236 |
lemma hypreal_epsilon_gt_zero: "0 < epsilon" |
237 |
by (simp add: hypreal_epsilon_inverse_omega) |
|
238 |
||
10751 | 239 |
end |