author | paulson |
Fri, 09 Jan 2004 10:46:18 +0100 | |
changeset 14348 | 744c868ee0b7 |
parent 14341 | a09441bd4f1e |
child 14361 | ad2f5da643b4 |
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 |
|
5 |
Description : Construction of hyperreals using ultrafilters |
|
13487 | 6 |
*) |
10751 | 7 |
|
14299 | 8 |
theory HyperDef = Filter + Real |
9 |
files ("fuf.ML"): (*Warning: file fuf.ML refers to the name Hyperdef!*) |
|
10751 | 10 |
|
11 |
||
12 |
constdefs |
|
14299 | 13 |
|
14 |
FreeUltrafilterNat :: "nat set set" ("\\<U>") |
|
15 |
"FreeUltrafilterNat == (SOME U. U \<in> FreeUltrafilter (UNIV:: nat set))" |
|
16 |
||
17 |
hyprel :: "((nat=>real)*(nat=>real)) set" |
|
18 |
"hyprel == {p. \<exists>X Y. p = ((X::nat=>real),Y) & |
|
10751 | 19 |
{n::nat. X(n) = Y(n)}: FreeUltrafilterNat}" |
20 |
||
14299 | 21 |
typedef hypreal = "UNIV//hyprel" |
22 |
by (auto simp add: quotient_def) |
|
10751 | 23 |
|
14299 | 24 |
instance hypreal :: ord .. |
25 |
instance hypreal :: zero .. |
|
26 |
instance hypreal :: one .. |
|
27 |
instance hypreal :: plus .. |
|
28 |
instance hypreal :: times .. |
|
29 |
instance hypreal :: minus .. |
|
30 |
instance hypreal :: inverse .. |
|
10751 | 31 |
|
32 |
||
14299 | 33 |
defs (overloaded) |
34 |
||
35 |
hypreal_zero_def: |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
36 |
"0 == Abs_hypreal(hyprel``{%n::nat. (0::real)})" |
10751 | 37 |
|
14299 | 38 |
hypreal_one_def: |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
39 |
"1 == Abs_hypreal(hyprel``{%n::nat. (1::real)})" |
10751 | 40 |
|
14299 | 41 |
hypreal_minus_def: |
42 |
"- P == Abs_hypreal(\<Union>X \<in> Rep_hypreal(P). hyprel``{%n::nat. - (X n)})" |
|
10751 | 43 |
|
14299 | 44 |
hypreal_diff_def: |
10751 | 45 |
"x - y == x + -(y::hypreal)" |
46 |
||
14299 | 47 |
hypreal_inverse_def: |
48 |
"inverse P == Abs_hypreal(\<Union>X \<in> Rep_hypreal(P). |
|
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11713
diff
changeset
|
49 |
hyprel``{%n. if X n = 0 then 0 else inverse (X n)})" |
10751 | 50 |
|
14299 | 51 |
hypreal_divide_def: |
10751 | 52 |
"P / Q::hypreal == P * inverse Q" |
13487 | 53 |
|
10751 | 54 |
constdefs |
55 |
||
14299 | 56 |
hypreal_of_real :: "real => hypreal" |
10834 | 57 |
"hypreal_of_real r == Abs_hypreal(hyprel``{%n::nat. r})" |
10751 | 58 |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
59 |
omega :: hypreal (*an infinite number = [<1,2,3,...>] *) |
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
60 |
"omega == Abs_hypreal(hyprel``{%n::nat. real (Suc n)})" |
13487 | 61 |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
62 |
epsilon :: hypreal (*an infinitesimal number = [<1,1/2,1/3,...>] *) |
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
63 |
"epsilon == Abs_hypreal(hyprel``{%n::nat. inverse (real (Suc n))})" |
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
64 |
|
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
65 |
syntax (xsymbols) |
14299 | 66 |
omega :: hypreal ("\<omega>") |
67 |
epsilon :: hypreal ("\<epsilon>") |
|
10919
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents:
10834
diff
changeset
|
68 |
|
10751 | 69 |
|
13487 | 70 |
defs |
71 |
||
14299 | 72 |
hypreal_add_def: |
73 |
"P + Q == Abs_hypreal(\<Union>X \<in> Rep_hypreal(P). \<Union>Y \<in> Rep_hypreal(Q). |
|
10834 | 74 |
hyprel``{%n::nat. X n + Y n})" |
10751 | 75 |
|
14299 | 76 |
hypreal_mult_def: |
77 |
"P * Q == Abs_hypreal(\<Union>X \<in> Rep_hypreal(P). \<Union>Y \<in> Rep_hypreal(Q). |
|
10834 | 78 |
hyprel``{%n::nat. X n * Y n})" |
10751 | 79 |
|
14299 | 80 |
hypreal_less_def: |
81 |
"P < (Q::hypreal) == \<exists>X Y. X \<in> Rep_hypreal(P) & |
|
82 |
Y \<in> Rep_hypreal(Q) & |
|
83 |
{n::nat. X n < Y n} \<in> FreeUltrafilterNat" |
|
84 |
hypreal_le_def: |
|
13487 | 85 |
"P <= (Q::hypreal) == ~(Q < P)" |
10751 | 86 |
|
14329 | 87 |
hrabs_def: "abs (r::hypreal) == (if 0 \<le> r then r else -r)" |
88 |
||
89 |
||
90 |
subsection{*The Set of Naturals is not Finite*} |
|
14299 | 91 |
|
92 |
(*** based on James' proof that the set of naturals is not finite ***) |
|
14329 | 93 |
lemma finite_exhausts [rule_format]: |
94 |
"finite (A::nat set) --> (\<exists>n. \<forall>m. Suc (n + m) \<notin> A)" |
|
14299 | 95 |
apply (rule impI) |
14301 | 96 |
apply (erule_tac F = A in finite_induct) |
97 |
apply (blast, erule exE) |
|
14299 | 98 |
apply (rule_tac x = "n + x" in exI) |
14301 | 99 |
apply (rule allI, erule_tac x = "x + m" in allE) |
14299 | 100 |
apply (auto simp add: add_ac) |
101 |
done |
|
102 |
||
14329 | 103 |
lemma finite_not_covers [rule_format (no_asm)]: |
104 |
"finite (A :: nat set) --> (\<exists>n. n \<notin>A)" |
|
14301 | 105 |
by (rule impI, drule finite_exhausts, blast) |
14299 | 106 |
|
107 |
lemma not_finite_nat: "~ finite(UNIV:: nat set)" |
|
14301 | 108 |
by (fast dest!: finite_exhausts) |
14299 | 109 |
|
14329 | 110 |
|
111 |
subsection{*Existence of Free Ultrafilter over the Naturals*} |
|
112 |
||
113 |
text{*Also, proof of various properties of @{term FreeUltrafilterNat}: |
|
114 |
an arbitrary free ultrafilter*} |
|
14299 | 115 |
|
116 |
lemma FreeUltrafilterNat_Ex: "\<exists>U. U: FreeUltrafilter (UNIV::nat set)" |
|
14301 | 117 |
by (rule not_finite_nat [THEN FreeUltrafilter_Ex]) |
14299 | 118 |
|
14301 | 119 |
lemma FreeUltrafilterNat_mem [simp]: |
14299 | 120 |
"FreeUltrafilterNat: FreeUltrafilter(UNIV:: nat set)" |
121 |
apply (unfold FreeUltrafilterNat_def) |
|
122 |
apply (rule FreeUltrafilterNat_Ex [THEN exE]) |
|
14301 | 123 |
apply (rule someI2, assumption+) |
14299 | 124 |
done |
125 |
||
126 |
lemma FreeUltrafilterNat_finite: "finite x ==> x \<notin> FreeUltrafilterNat" |
|
127 |
apply (unfold FreeUltrafilterNat_def) |
|
128 |
apply (rule FreeUltrafilterNat_Ex [THEN exE]) |
|
14301 | 129 |
apply (rule someI2, assumption) |
14299 | 130 |
apply (blast dest: mem_FreeUltrafiltersetD1) |
131 |
done |
|
132 |
||
133 |
lemma FreeUltrafilterNat_not_finite: "x: FreeUltrafilterNat ==> ~ finite x" |
|
14301 | 134 |
by (blast dest: FreeUltrafilterNat_finite) |
14299 | 135 |
|
14301 | 136 |
lemma FreeUltrafilterNat_empty [simp]: "{} \<notin> FreeUltrafilterNat" |
14299 | 137 |
apply (unfold FreeUltrafilterNat_def) |
138 |
apply (rule FreeUltrafilterNat_Ex [THEN exE]) |
|
14301 | 139 |
apply (rule someI2, assumption) |
140 |
apply (blast dest: FreeUltrafilter_Ultrafilter Ultrafilter_Filter |
|
141 |
Filter_empty_not_mem) |
|
14299 | 142 |
done |
143 |
||
14329 | 144 |
lemma FreeUltrafilterNat_Int: |
145 |
"[| X: FreeUltrafilterNat; Y: FreeUltrafilterNat |] |
|
14299 | 146 |
==> X Int Y \<in> FreeUltrafilterNat" |
147 |
apply (cut_tac FreeUltrafilterNat_mem) |
|
148 |
apply (blast dest: FreeUltrafilter_Ultrafilter Ultrafilter_Filter mem_FiltersetD1) |
|
149 |
done |
|
150 |
||
14329 | 151 |
lemma FreeUltrafilterNat_subset: |
152 |
"[| X: FreeUltrafilterNat; X <= Y |] |
|
14299 | 153 |
==> Y \<in> FreeUltrafilterNat" |
154 |
apply (cut_tac FreeUltrafilterNat_mem) |
|
155 |
apply (blast dest: FreeUltrafilter_Ultrafilter Ultrafilter_Filter mem_FiltersetD2) |
|
156 |
done |
|
157 |
||
14329 | 158 |
lemma FreeUltrafilterNat_Compl: |
159 |
"X: FreeUltrafilterNat ==> -X \<notin> FreeUltrafilterNat" |
|
14301 | 160 |
apply safe |
161 |
apply (drule FreeUltrafilterNat_Int, assumption, auto) |
|
14299 | 162 |
done |
163 |
||
14329 | 164 |
lemma FreeUltrafilterNat_Compl_mem: |
165 |
"X\<notin> FreeUltrafilterNat ==> -X \<in> FreeUltrafilterNat" |
|
14299 | 166 |
apply (cut_tac FreeUltrafilterNat_mem [THEN FreeUltrafilter_iff [THEN iffD1]]) |
14301 | 167 |
apply (safe, drule_tac x = X in bspec) |
14299 | 168 |
apply (auto simp add: UNIV_diff_Compl) |
169 |
done |
|
170 |
||
14329 | 171 |
lemma FreeUltrafilterNat_Compl_iff1: |
172 |
"(X \<notin> FreeUltrafilterNat) = (-X: FreeUltrafilterNat)" |
|
14301 | 173 |
by (blast dest: FreeUltrafilterNat_Compl FreeUltrafilterNat_Compl_mem) |
14299 | 174 |
|
14329 | 175 |
lemma FreeUltrafilterNat_Compl_iff2: |
176 |
"(X: FreeUltrafilterNat) = (-X \<notin> FreeUltrafilterNat)" |
|
14301 | 177 |
by (auto simp add: FreeUltrafilterNat_Compl_iff1 [symmetric]) |
14299 | 178 |
|
14301 | 179 |
lemma FreeUltrafilterNat_UNIV [simp]: "(UNIV::nat set) \<in> FreeUltrafilterNat" |
180 |
by (rule FreeUltrafilterNat_mem [THEN FreeUltrafilter_Ultrafilter, THEN Ultrafilter_Filter, THEN mem_FiltersetD4]) |
|
14299 | 181 |
|
14301 | 182 |
lemma FreeUltrafilterNat_Nat_set [simp]: "UNIV \<in> FreeUltrafilterNat" |
183 |
by auto |
|
14299 | 184 |
|
14329 | 185 |
lemma FreeUltrafilterNat_Nat_set_refl [intro]: |
186 |
"{n. P(n) = P(n)} \<in> FreeUltrafilterNat" |
|
14301 | 187 |
by simp |
14299 | 188 |
|
189 |
lemma FreeUltrafilterNat_P: "{n::nat. P} \<in> FreeUltrafilterNat ==> P" |
|
14301 | 190 |
by (rule ccontr, simp) |
14299 | 191 |
|
192 |
lemma FreeUltrafilterNat_Ex_P: "{n. P(n)} \<in> FreeUltrafilterNat ==> \<exists>n. P(n)" |
|
14301 | 193 |
by (rule ccontr, simp) |
14299 | 194 |
|
195 |
lemma FreeUltrafilterNat_all: "\<forall>n. P(n) ==> {n. P(n)} \<in> FreeUltrafilterNat" |
|
14301 | 196 |
by (auto intro: FreeUltrafilterNat_Nat_set) |
14299 | 197 |
|
14329 | 198 |
|
199 |
text{*Define and use Ultrafilter tactics*} |
|
14299 | 200 |
use "fuf.ML" |
201 |
||
202 |
method_setup fuf = {* |
|
203 |
Method.ctxt_args (fn ctxt => |
|
204 |
Method.METHOD (fn facts => |
|
205 |
fuf_tac (Classical.get_local_claset ctxt, |
|
206 |
Simplifier.get_local_simpset ctxt) 1)) *} |
|
207 |
"free ultrafilter tactic" |
|
208 |
||
209 |
method_setup ultra = {* |
|
210 |
Method.ctxt_args (fn ctxt => |
|
211 |
Method.METHOD (fn facts => |
|
212 |
ultra_tac (Classical.get_local_claset ctxt, |
|
213 |
Simplifier.get_local_simpset ctxt) 1)) *} |
|
214 |
"ultrafilter tactic" |
|
215 |
||
216 |
||
14329 | 217 |
text{*One further property of our free ultrafilter*} |
218 |
lemma FreeUltrafilterNat_Un: |
|
219 |
"X Un Y: FreeUltrafilterNat |
|
14299 | 220 |
==> X: FreeUltrafilterNat | Y: FreeUltrafilterNat" |
221 |
apply auto |
|
14301 | 222 |
apply ultra |
14299 | 223 |
done |
224 |
||
225 |
||
14329 | 226 |
subsection{*Properties of @{term hyprel}*} |
227 |
||
228 |
text{*Proving that @{term hyprel} is an equivalence relation*} |
|
14299 | 229 |
|
230 |
lemma hyprel_iff: "((X,Y): hyprel) = ({n. X n = Y n}: FreeUltrafilterNat)" |
|
14301 | 231 |
by (unfold hyprel_def, fast) |
14299 | 232 |
|
233 |
lemma hyprel_refl: "(x,x): hyprel" |
|
234 |
apply (unfold hyprel_def) |
|
235 |
apply (auto simp add: FreeUltrafilterNat_Nat_set) |
|
236 |
done |
|
237 |
||
238 |
lemma hyprel_sym [rule_format (no_asm)]: "(x,y): hyprel --> (y,x):hyprel" |
|
14301 | 239 |
by (simp add: hyprel_def eq_commute) |
14299 | 240 |
|
241 |
lemma hyprel_trans: |
|
242 |
"[|(x,y): hyprel; (y,z):hyprel|] ==> (x,z):hyprel" |
|
14301 | 243 |
apply (unfold hyprel_def, auto, ultra) |
14299 | 244 |
done |
245 |
||
246 |
lemma equiv_hyprel: "equiv UNIV hyprel" |
|
247 |
apply (simp add: equiv_def refl_def sym_def trans_def hyprel_refl) |
|
248 |
apply (blast intro: hyprel_sym hyprel_trans) |
|
249 |
done |
|
250 |
||
251 |
(* (hyprel `` {x} = hyprel `` {y}) = ((x,y) \<in> hyprel) *) |
|
252 |
lemmas equiv_hyprel_iff = |
|
253 |
eq_equiv_class_iff [OF equiv_hyprel UNIV_I UNIV_I, simp] |
|
254 |
||
14301 | 255 |
lemma hyprel_in_hypreal [simp]: "hyprel``{x}:hypreal" |
256 |
by (unfold hypreal_def hyprel_def quotient_def, blast) |
|
14299 | 257 |
|
258 |
lemma inj_on_Abs_hypreal: "inj_on Abs_hypreal hypreal" |
|
259 |
apply (rule inj_on_inverseI) |
|
260 |
apply (erule Abs_hypreal_inverse) |
|
261 |
done |
|
262 |
||
263 |
declare inj_on_Abs_hypreal [THEN inj_on_iff, simp] |
|
14301 | 264 |
Abs_hypreal_inverse [simp] |
14299 | 265 |
|
266 |
declare equiv_hyprel [THEN eq_equiv_class_iff, simp] |
|
267 |
||
268 |
declare hyprel_iff [iff] |
|
269 |
||
270 |
lemmas eq_hyprelD = eq_equiv_class [OF _ equiv_hyprel] |
|
271 |
||
272 |
lemma inj_Rep_hypreal: "inj(Rep_hypreal)" |
|
273 |
apply (rule inj_on_inverseI) |
|
274 |
apply (rule Rep_hypreal_inverse) |
|
275 |
done |
|
276 |
||
14301 | 277 |
lemma lemma_hyprel_refl [simp]: "x \<in> hyprel `` {x}" |
278 |
apply (unfold hyprel_def, safe) |
|
14299 | 279 |
apply (auto intro!: FreeUltrafilterNat_Nat_set) |
280 |
done |
|
281 |
||
14301 | 282 |
lemma hypreal_empty_not_mem [simp]: "{} \<notin> hypreal" |
14299 | 283 |
apply (unfold hypreal_def) |
284 |
apply (auto elim!: quotientE equalityCE) |
|
285 |
done |
|
286 |
||
14301 | 287 |
lemma Rep_hypreal_nonempty [simp]: "Rep_hypreal x \<noteq> {}" |
288 |
by (cut_tac x = x in Rep_hypreal, auto) |
|
14299 | 289 |
|
290 |
||
14329 | 291 |
subsection{*@{term hypreal_of_real}: |
292 |
the Injection from @{typ real} to @{typ hypreal}*} |
|
14299 | 293 |
|
294 |
lemma inj_hypreal_of_real: "inj(hypreal_of_real)" |
|
295 |
apply (rule inj_onI) |
|
296 |
apply (unfold hypreal_of_real_def) |
|
297 |
apply (drule inj_on_Abs_hypreal [THEN inj_onD]) |
|
298 |
apply (rule hyprel_in_hypreal)+ |
|
299 |
apply (drule eq_equiv_class) |
|
300 |
apply (rule equiv_hyprel) |
|
301 |
apply (simp_all add: split: split_if_asm) |
|
302 |
done |
|
303 |
||
304 |
lemma eq_Abs_hypreal: |
|
305 |
"(!!x y. z = Abs_hypreal(hyprel``{x}) ==> P) ==> P" |
|
306 |
apply (rule_tac x1=z in Rep_hypreal [unfolded hypreal_def, THEN quotientE]) |
|
14301 | 307 |
apply (drule_tac f = Abs_hypreal in arg_cong) |
14299 | 308 |
apply (force simp add: Rep_hypreal_inverse) |
309 |
done |
|
310 |
||
14329 | 311 |
|
312 |
subsection{*Hyperreal Addition*} |
|
313 |
||
314 |
lemma hypreal_add_congruent2: |
|
315 |
"congruent2 hyprel (%X Y. hyprel``{%n. X n + Y n})" |
|
316 |
apply (unfold congruent2_def, auto, ultra) |
|
317 |
done |
|
318 |
||
319 |
lemma hypreal_add: |
|
320 |
"Abs_hypreal(hyprel``{%n. X n}) + Abs_hypreal(hyprel``{%n. Y n}) = |
|
321 |
Abs_hypreal(hyprel``{%n. X n + Y n})" |
|
322 |
apply (unfold hypreal_add_def) |
|
323 |
apply (simp add: UN_equiv_class2 [OF equiv_hyprel hypreal_add_congruent2]) |
|
324 |
done |
|
325 |
||
326 |
lemma hypreal_add_commute: "(z::hypreal) + w = w + z" |
|
327 |
apply (rule_tac z = z in eq_Abs_hypreal) |
|
328 |
apply (rule_tac z = w in eq_Abs_hypreal) |
|
14334 | 329 |
apply (simp add: add_ac hypreal_add) |
14329 | 330 |
done |
331 |
||
332 |
lemma hypreal_add_assoc: "((z1::hypreal) + z2) + z3 = z1 + (z2 + z3)" |
|
333 |
apply (rule_tac z = z1 in eq_Abs_hypreal) |
|
334 |
apply (rule_tac z = z2 in eq_Abs_hypreal) |
|
335 |
apply (rule_tac z = z3 in eq_Abs_hypreal) |
|
336 |
apply (simp add: hypreal_add real_add_assoc) |
|
337 |
done |
|
338 |
||
14331 | 339 |
lemma hypreal_add_zero_left: "(0::hypreal) + z = z" |
14329 | 340 |
apply (unfold hypreal_zero_def) |
341 |
apply (rule_tac z = z in eq_Abs_hypreal) |
|
342 |
apply (simp add: hypreal_add) |
|
343 |
done |
|
344 |
||
345 |
instance hypreal :: plus_ac0 |
|
346 |
by (intro_classes, |
|
347 |
(assumption | |
|
348 |
rule hypreal_add_commute hypreal_add_assoc hypreal_add_zero_left)+) |
|
349 |
||
350 |
lemma hypreal_add_zero_right [simp]: "z + (0::hypreal) = z" |
|
351 |
by (simp add: hypreal_add_zero_left hypreal_add_commute) |
|
352 |
||
353 |
||
354 |
subsection{*Additive inverse on @{typ hypreal}*} |
|
14299 | 355 |
|
356 |
lemma hypreal_minus_congruent: |
|
357 |
"congruent hyprel (%X. hyprel``{%n. - (X n)})" |
|
358 |
by (force simp add: congruent_def) |
|
359 |
||
360 |
lemma hypreal_minus: |
|
361 |
"- (Abs_hypreal(hyprel``{%n. X n})) = Abs_hypreal(hyprel `` {%n. -(X n)})" |
|
362 |
apply (unfold hypreal_minus_def) |
|
14301 | 363 |
apply (rule_tac f = Abs_hypreal in arg_cong) |
364 |
apply (simp add: hyprel_in_hypreal [THEN Abs_hypreal_inverse] |
|
14299 | 365 |
UN_equiv_class [OF equiv_hyprel hypreal_minus_congruent]) |
366 |
done |
|
367 |
||
14329 | 368 |
lemma hypreal_diff: |
369 |
"Abs_hypreal(hyprel``{%n. X n}) - Abs_hypreal(hyprel``{%n. Y n}) = |
|
14299 | 370 |
Abs_hypreal(hyprel``{%n. X n - Y n})" |
14301 | 371 |
apply (simp add: hypreal_diff_def hypreal_minus hypreal_add) |
14299 | 372 |
done |
373 |
||
14301 | 374 |
lemma hypreal_add_minus [simp]: "z + -z = (0::hypreal)" |
14299 | 375 |
apply (unfold hypreal_zero_def) |
14301 | 376 |
apply (rule_tac z = z in eq_Abs_hypreal) |
14299 | 377 |
apply (simp add: hypreal_minus hypreal_add) |
378 |
done |
|
379 |
||
14331 | 380 |
lemma hypreal_add_minus_left: "-z + z = (0::hypreal)" |
14301 | 381 |
by (simp add: hypreal_add_commute hypreal_add_minus) |
14299 | 382 |
|
14329 | 383 |
|
384 |
subsection{*Hyperreal Multiplication*} |
|
14299 | 385 |
|
386 |
lemma hypreal_mult_congruent2: |
|
387 |
"congruent2 hyprel (%X Y. hyprel``{%n. X n * Y n})" |
|
14301 | 388 |
apply (unfold congruent2_def, auto, ultra) |
14299 | 389 |
done |
390 |
||
391 |
lemma hypreal_mult: |
|
392 |
"Abs_hypreal(hyprel``{%n. X n}) * Abs_hypreal(hyprel``{%n. Y n}) = |
|
393 |
Abs_hypreal(hyprel``{%n. X n * Y n})" |
|
394 |
apply (unfold hypreal_mult_def) |
|
395 |
apply (simp add: UN_equiv_class2 [OF equiv_hyprel hypreal_mult_congruent2]) |
|
396 |
done |
|
397 |
||
398 |
lemma hypreal_mult_commute: "(z::hypreal) * w = w * z" |
|
14301 | 399 |
apply (rule_tac z = z in eq_Abs_hypreal) |
400 |
apply (rule_tac z = w in eq_Abs_hypreal) |
|
14331 | 401 |
apply (simp add: hypreal_mult mult_ac) |
14299 | 402 |
done |
403 |
||
404 |
lemma hypreal_mult_assoc: "((z1::hypreal) * z2) * z3 = z1 * (z2 * z3)" |
|
14301 | 405 |
apply (rule_tac z = z1 in eq_Abs_hypreal) |
406 |
apply (rule_tac z = z2 in eq_Abs_hypreal) |
|
407 |
apply (rule_tac z = z3 in eq_Abs_hypreal) |
|
14331 | 408 |
apply (simp add: hypreal_mult mult_assoc) |
14299 | 409 |
done |
410 |
||
14331 | 411 |
lemma hypreal_mult_1: "(1::hypreal) * z = z" |
14299 | 412 |
apply (unfold hypreal_one_def) |
14301 | 413 |
apply (rule_tac z = z in eq_Abs_hypreal) |
14299 | 414 |
apply (simp add: hypreal_mult) |
415 |
done |
|
14301 | 416 |
|
14329 | 417 |
lemma hypreal_add_mult_distrib: |
418 |
"((z1::hypreal) + z2) * w = (z1 * w) + (z2 * w)" |
|
14301 | 419 |
apply (rule_tac z = z1 in eq_Abs_hypreal) |
420 |
apply (rule_tac z = z2 in eq_Abs_hypreal) |
|
421 |
apply (rule_tac z = w in eq_Abs_hypreal) |
|
14334 | 422 |
apply (simp add: hypreal_mult hypreal_add left_distrib) |
14299 | 423 |
done |
424 |
||
14331 | 425 |
text{*one and zero are distinct*} |
14299 | 426 |
lemma hypreal_zero_not_eq_one: "0 \<noteq> (1::hypreal)" |
427 |
apply (unfold hypreal_zero_def hypreal_one_def) |
|
428 |
apply (auto simp add: real_zero_not_eq_one) |
|
429 |
done |
|
430 |
||
431 |
||
14329 | 432 |
subsection{*Multiplicative Inverse on @{typ hypreal} *} |
14299 | 433 |
|
434 |
lemma hypreal_inverse_congruent: |
|
435 |
"congruent hyprel (%X. hyprel``{%n. if X n = 0 then 0 else inverse(X n)})" |
|
436 |
apply (unfold congruent_def) |
|
14301 | 437 |
apply (auto, ultra) |
14299 | 438 |
done |
439 |
||
440 |
lemma hypreal_inverse: |
|
441 |
"inverse (Abs_hypreal(hyprel``{%n. X n})) = |
|
442 |
Abs_hypreal(hyprel `` {%n. if X n = 0 then 0 else inverse(X n)})" |
|
443 |
apply (unfold hypreal_inverse_def) |
|
14301 | 444 |
apply (rule_tac f = Abs_hypreal in arg_cong) |
445 |
apply (simp add: hyprel_in_hypreal [THEN Abs_hypreal_inverse] |
|
14299 | 446 |
UN_equiv_class [OF equiv_hyprel hypreal_inverse_congruent]) |
447 |
done |
|
448 |
||
14331 | 449 |
lemma hypreal_mult_inverse: |
14299 | 450 |
"x \<noteq> 0 ==> x*inverse(x) = (1::hypreal)" |
451 |
apply (unfold hypreal_one_def hypreal_zero_def) |
|
14301 | 452 |
apply (rule_tac z = x in eq_Abs_hypreal) |
14299 | 453 |
apply (simp add: hypreal_inverse hypreal_mult) |
454 |
apply (drule FreeUltrafilterNat_Compl_mem) |
|
14334 | 455 |
apply (blast intro!: right_inverse FreeUltrafilterNat_subset) |
14299 | 456 |
done |
457 |
||
14331 | 458 |
lemma hypreal_mult_inverse_left: |
14329 | 459 |
"x \<noteq> 0 ==> inverse(x)*x = (1::hypreal)" |
14301 | 460 |
by (simp add: hypreal_mult_inverse hypreal_mult_commute) |
14299 | 461 |
|
14331 | 462 |
instance hypreal :: field |
463 |
proof |
|
464 |
fix x y z :: hypreal |
|
465 |
show "(x + y) + z = x + (y + z)" by (rule hypreal_add_assoc) |
|
466 |
show "x + y = y + x" by (rule hypreal_add_commute) |
|
467 |
show "0 + x = x" by simp |
|
468 |
show "- x + x = 0" by (simp add: hypreal_add_minus_left) |
|
469 |
show "x - y = x + (-y)" by (simp add: hypreal_diff_def) |
|
470 |
show "(x * y) * z = x * (y * z)" by (rule hypreal_mult_assoc) |
|
471 |
show "x * y = y * x" by (rule hypreal_mult_commute) |
|
472 |
show "1 * x = x" by (simp add: hypreal_mult_1) |
|
473 |
show "(x + y) * z = x * z + y * z" by (simp add: hypreal_add_mult_distrib) |
|
474 |
show "0 \<noteq> (1::hypreal)" by (rule hypreal_zero_not_eq_one) |
|
475 |
show "x \<noteq> 0 ==> inverse x * x = 1" by (simp add: hypreal_mult_inverse_left) |
|
476 |
show "y \<noteq> 0 ==> x / y = x * inverse y" by (simp add: hypreal_divide_def) |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14334
diff
changeset
|
477 |
assume eq: "z+x = z+y" |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14334
diff
changeset
|
478 |
hence "(-z + z) + x = (-z + z) + y" by (simp only: eq hypreal_add_assoc) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14334
diff
changeset
|
479 |
thus "x = y" by (simp add: hypreal_add_minus_left) |
14331 | 480 |
qed |
481 |
||
482 |
||
483 |
lemma HYPREAL_INVERSE_ZERO: "inverse 0 = (0::hypreal)" |
|
484 |
by (simp add: hypreal_inverse hypreal_zero_def) |
|
485 |
||
486 |
lemma HYPREAL_DIVISION_BY_ZERO: "a / (0::hypreal) = 0" |
|
487 |
by (simp add: hypreal_divide_def HYPREAL_INVERSE_ZERO |
|
488 |
hypreal_mult_commute [of a]) |
|
489 |
||
490 |
instance hypreal :: division_by_zero |
|
491 |
proof |
|
492 |
fix x :: hypreal |
|
493 |
show "inverse 0 = (0::hypreal)" by (rule HYPREAL_INVERSE_ZERO) |
|
494 |
show "x/0 = 0" by (rule HYPREAL_DIVISION_BY_ZERO) |
|
495 |
qed |
|
496 |
||
14329 | 497 |
|
498 |
subsection{*Theorems for Ordering*} |
|
499 |
||
500 |
text{*TODO: define @{text "\<le>"} as the primitive concept and quickly |
|
501 |
establish membership in class @{text linorder}. Then proofs could be |
|
502 |
simplified, since properties of @{text "<"} would be generic.*} |
|
14299 | 503 |
|
14329 | 504 |
text{*TODO: The following theorem should be used througout the proofs |
505 |
as it probably makes many of them more straightforward.*} |
|
506 |
lemma hypreal_less: |
|
507 |
"(Abs_hypreal(hyprel``{%n. X n}) < Abs_hypreal(hyprel``{%n. Y n})) = |
|
508 |
({n. X n < Y n} \<in> FreeUltrafilterNat)" |
|
509 |
apply (unfold hypreal_less_def) |
|
510 |
apply (auto intro!: lemma_hyprel_refl, ultra) |
|
14299 | 511 |
done |
512 |
||
513 |
lemma hypreal_less_not_refl: "~ (R::hypreal) < R" |
|
14301 | 514 |
apply (rule_tac z = R in eq_Abs_hypreal) |
515 |
apply (auto simp add: hypreal_less_def, ultra) |
|
14299 | 516 |
done |
517 |
||
518 |
lemmas hypreal_less_irrefl = hypreal_less_not_refl [THEN notE, standard] |
|
519 |
declare hypreal_less_irrefl [elim!] |
|
520 |
||
521 |
lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \<noteq> y" |
|
14301 | 522 |
by (auto simp add: hypreal_less_not_refl) |
14299 | 523 |
|
524 |
lemma hypreal_less_trans: "!!(R1::hypreal). [| R1 < R2; R2 < R3 |] ==> R1 < R3" |
|
14301 | 525 |
apply (rule_tac z = R1 in eq_Abs_hypreal) |
526 |
apply (rule_tac z = R2 in eq_Abs_hypreal) |
|
527 |
apply (rule_tac z = R3 in eq_Abs_hypreal) |
|
528 |
apply (auto intro!: exI simp add: hypreal_less_def, ultra) |
|
14299 | 529 |
done |
530 |
||
531 |
lemma hypreal_less_asym: "!! (R1::hypreal). [| R1 < R2; R2 < R1 |] ==> P" |
|
14301 | 532 |
apply (drule hypreal_less_trans, assumption) |
14299 | 533 |
apply (simp add: hypreal_less_not_refl) |
534 |
done |
|
535 |
||
536 |
||
14329 | 537 |
subsection{*Trichotomy: the hyperreals are Linearly Ordered*} |
14299 | 538 |
|
539 |
lemma lemma_hyprel_0_mem: "\<exists>x. x: hyprel `` {%n. 0}" |
|
540 |
apply (unfold hyprel_def) |
|
14301 | 541 |
apply (rule_tac x = "%n. 0" in exI, safe) |
14299 | 542 |
apply (auto intro!: FreeUltrafilterNat_Nat_set) |
543 |
done |
|
544 |
||
545 |
lemma hypreal_trichotomy: "0 < x | x = 0 | x < (0::hypreal)" |
|
546 |
apply (unfold hypreal_zero_def) |
|
14301 | 547 |
apply (rule_tac z = x in eq_Abs_hypreal) |
14299 | 548 |
apply (auto simp add: hypreal_less_def) |
14301 | 549 |
apply (cut_tac lemma_hyprel_0_mem, erule exE) |
550 |
apply (drule_tac x = xa in spec) |
|
551 |
apply (drule_tac x = x in spec) |
|
552 |
apply (cut_tac x = x in lemma_hyprel_refl, auto) |
|
553 |
apply (drule_tac x = x in spec) |
|
554 |
apply (drule_tac x = xa in spec, auto, ultra) |
|
14299 | 555 |
done |
556 |
||
557 |
lemma hypreal_trichotomyE: |
|
558 |
"[| (0::hypreal) < x ==> P; |
|
559 |
x = 0 ==> P; |
|
560 |
x < 0 ==> P |] ==> P" |
|
14301 | 561 |
apply (insert hypreal_trichotomy [of x], blast) |
14299 | 562 |
done |
563 |
||
564 |
lemma hypreal_less_minus_iff: "((x::hypreal) < y) = (0 < y + -x)" |
|
14301 | 565 |
apply (rule_tac z = x in eq_Abs_hypreal) |
566 |
apply (rule_tac z = y in eq_Abs_hypreal) |
|
14299 | 567 |
apply (auto simp add: hypreal_add hypreal_zero_def hypreal_minus hypreal_less) |
568 |
done |
|
569 |
||
570 |
lemma hypreal_less_minus_iff2: "((x::hypreal) < y) = (x + -y < 0)" |
|
14301 | 571 |
apply (rule_tac z = x in eq_Abs_hypreal) |
572 |
apply (rule_tac z = y in eq_Abs_hypreal) |
|
14299 | 573 |
apply (auto simp add: hypreal_add hypreal_zero_def hypreal_minus hypreal_less) |
574 |
done |
|
575 |
||
576 |
lemma hypreal_eq_minus_iff2: "((x::hypreal) = y) = (0 = y + - x)" |
|
577 |
apply auto |
|
14331 | 578 |
apply (rule Ring_and_Field.add_right_cancel [of _ "-x", THEN iffD1], auto) |
14299 | 579 |
done |
580 |
||
581 |
lemma hypreal_linear: "(x::hypreal) < y | x = y | y < x" |
|
582 |
apply (subst hypreal_eq_minus_iff2) |
|
14301 | 583 |
apply (rule_tac x1 = x in hypreal_less_minus_iff [THEN ssubst]) |
584 |
apply (rule_tac x1 = y in hypreal_less_minus_iff2 [THEN ssubst]) |
|
585 |
apply (rule hypreal_trichotomyE, auto) |
|
14299 | 586 |
done |
587 |
||
588 |
lemma hypreal_neq_iff: "((w::hypreal) \<noteq> z) = (w<z | z<w)" |
|
14301 | 589 |
by (cut_tac hypreal_linear, blast) |
14299 | 590 |
|
591 |
lemma hypreal_linear_less2: "!!(x::hypreal). [| x < y ==> P; x = y ==> P; |
|
592 |
y < x ==> P |] ==> P" |
|
14301 | 593 |
apply (cut_tac x = x and y = y in hypreal_linear, auto) |
14299 | 594 |
done |
595 |
||
14329 | 596 |
|
597 |
subsection{*Properties of The @{text "\<le>"} Relation*} |
|
14299 | 598 |
|
599 |
lemma hypreal_le: |
|
600 |
"(Abs_hypreal(hyprel``{%n. X n}) <= |
|
601 |
Abs_hypreal(hyprel``{%n. Y n})) = |
|
602 |
({n. X n <= Y n} \<in> FreeUltrafilterNat)" |
|
603 |
apply (unfold hypreal_le_def real_le_def) |
|
604 |
apply (auto simp add: hypreal_less) |
|
605 |
apply (ultra+) |
|
606 |
done |
|
607 |
||
608 |
lemma hypreal_le_imp_less_or_eq: "!!(x::hypreal). x <= y ==> x < y | x = y" |
|
609 |
apply (unfold hypreal_le_def) |
|
610 |
apply (cut_tac hypreal_linear) |
|
611 |
apply (fast elim: hypreal_less_irrefl hypreal_less_asym) |
|
612 |
done |
|
613 |
||
614 |
lemma hypreal_less_or_eq_imp_le: "z<w | z=w ==> z <=(w::hypreal)" |
|
615 |
apply (unfold hypreal_le_def) |
|
616 |
apply (cut_tac hypreal_linear) |
|
617 |
apply (fast elim: hypreal_less_irrefl hypreal_less_asym) |
|
618 |
done |
|
619 |
||
620 |
lemma hypreal_le_eq_less_or_eq: "(x <= (y::hypreal)) = (x < y | x=y)" |
|
621 |
by (blast intro!: hypreal_less_or_eq_imp_le dest: hypreal_le_imp_less_or_eq) |
|
622 |
||
623 |
lemmas hypreal_le_less = hypreal_le_eq_less_or_eq |
|
624 |
||
625 |
lemma hypreal_le_refl: "w <= (w::hypreal)" |
|
14301 | 626 |
by (simp add: hypreal_le_eq_less_or_eq) |
14299 | 627 |
|
628 |
(* Axiom 'linorder_linear' of class 'linorder': *) |
|
629 |
lemma hypreal_le_linear: "(z::hypreal) <= w | w <= z" |
|
14301 | 630 |
apply (simp add: hypreal_le_less) |
631 |
apply (cut_tac hypreal_linear, blast) |
|
14299 | 632 |
done |
633 |
||
634 |
lemma hypreal_le_trans: "[| i <= j; j <= k |] ==> i <= (k::hypreal)" |
|
635 |
apply (drule hypreal_le_imp_less_or_eq) |
|
636 |
apply (drule hypreal_le_imp_less_or_eq) |
|
637 |
apply (rule hypreal_less_or_eq_imp_le) |
|
638 |
apply (blast intro: hypreal_less_trans) |
|
639 |
done |
|
640 |
||
641 |
lemma hypreal_le_anti_sym: "[| z <= w; w <= z |] ==> z = (w::hypreal)" |
|
642 |
apply (drule hypreal_le_imp_less_or_eq) |
|
643 |
apply (drule hypreal_le_imp_less_or_eq) |
|
644 |
apply (fast elim: hypreal_less_irrefl hypreal_less_asym) |
|
645 |
done |
|
646 |
||
647 |
(* Axiom 'order_less_le' of class 'order': *) |
|
648 |
lemma hypreal_less_le: "((w::hypreal) < z) = (w <= z & w \<noteq> z)" |
|
14301 | 649 |
apply (simp add: hypreal_le_def hypreal_neq_iff) |
14299 | 650 |
apply (blast intro: hypreal_less_asym) |
651 |
done |
|
652 |
||
14329 | 653 |
instance hypreal :: order |
654 |
by (intro_classes, |
|
655 |
(assumption | |
|
656 |
rule hypreal_le_refl hypreal_le_trans hypreal_le_anti_sym |
|
657 |
hypreal_less_le)+) |
|
658 |
||
659 |
instance hypreal :: linorder |
|
660 |
by (intro_classes, rule hypreal_le_linear) |
|
661 |
||
662 |
||
663 |
lemma hypreal_add_less_mono1: "(A::hypreal) < B ==> A + C < B + C" |
|
664 |
apply (rule_tac z = A in eq_Abs_hypreal) |
|
665 |
apply (rule_tac z = B in eq_Abs_hypreal) |
|
666 |
apply (rule_tac z = C in eq_Abs_hypreal) |
|
667 |
apply (auto intro!: exI simp add: hypreal_less_def hypreal_add, ultra) |
|
668 |
done |
|
669 |
||
670 |
lemma hypreal_mult_order: "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y" |
|
671 |
apply (unfold hypreal_zero_def) |
|
672 |
apply (rule_tac z = x in eq_Abs_hypreal) |
|
673 |
apply (rule_tac z = y in eq_Abs_hypreal) |
|
674 |
apply (auto intro!: exI simp add: hypreal_less_def hypreal_mult, ultra) |
|
675 |
apply (auto intro: real_mult_order) |
|
676 |
done |
|
677 |
||
678 |
lemma hypreal_add_left_le_mono1: "(q1::hypreal) \<le> q2 ==> x + q1 \<le> x + q2" |
|
679 |
apply (drule order_le_imp_less_or_eq) |
|
680 |
apply (auto intro: order_less_imp_le hypreal_add_less_mono1 simp add: hypreal_add_commute) |
|
681 |
done |
|
682 |
||
683 |
lemma hypreal_mult_less_mono1: "[| (0::hypreal) < z; x < y |] ==> x*z < y*z" |
|
684 |
apply (rotate_tac 1) |
|
685 |
apply (drule hypreal_less_minus_iff [THEN iffD1]) |
|
686 |
apply (rule hypreal_less_minus_iff [THEN iffD2]) |
|
687 |
apply (drule hypreal_mult_order, assumption) |
|
14331 | 688 |
apply (simp add: right_distrib hypreal_mult_commute) |
14329 | 689 |
done |
690 |
||
691 |
lemma hypreal_mult_less_mono2: "[| (0::hypreal)<z; x<y |] ==> z*x<z*y" |
|
692 |
apply (simp (no_asm_simp) add: hypreal_mult_commute hypreal_mult_less_mono1) |
|
693 |
done |
|
694 |
||
695 |
subsection{*The Hyperreals Form an Ordered Field*} |
|
696 |
||
697 |
instance hypreal :: ordered_field |
|
698 |
proof |
|
699 |
fix x y z :: hypreal |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
700 |
show "0 < (1::hypreal)" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
701 |
by (unfold hypreal_one_def hypreal_zero_def hypreal_less_def, force) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
702 |
show "x \<le> y ==> z + x \<le> z + y" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
703 |
by (rule hypreal_add_left_le_mono1) |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
704 |
show "x < y ==> 0 < z ==> z * x < z * y" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
705 |
by (simp add: hypreal_mult_less_mono2) |
14329 | 706 |
show "\<bar>x\<bar> = (if x < 0 then -x else x)" |
707 |
by (auto dest: order_le_less_trans simp add: hrabs_def linorder_not_le) |
|
708 |
qed |
|
709 |
||
14331 | 710 |
lemma hypreal_mult_1_right: "z * (1::hypreal) = z" |
711 |
by (rule Ring_and_Field.mult_1_right) |
|
712 |
||
713 |
lemma hypreal_mult_minus_1 [simp]: "(- (1::hypreal)) * z = -z" |
|
714 |
by (simp) |
|
715 |
||
716 |
lemma hypreal_mult_minus_1_right [simp]: "z * (- (1::hypreal)) = -z" |
|
717 |
by (subst hypreal_mult_commute, simp) |
|
14329 | 718 |
|
719 |
(*Used ONCE: in NSA.ML*) |
|
720 |
lemma hypreal_minus_distrib1: "-(y + -(x::hypreal)) = x + -y" |
|
721 |
by (simp add: hypreal_add_commute) |
|
722 |
||
723 |
(*Used ONCE: in Lim.ML*) |
|
724 |
lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))" |
|
725 |
by (auto simp add: hypreal_add_assoc) |
|
726 |
||
14331 | 727 |
lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)" |
728 |
apply auto |
|
729 |
apply (rule Ring_and_Field.add_right_cancel [of _ "-y", THEN iffD1], auto) |
|
730 |
done |
|
731 |
||
732 |
(*Used 3 TIMES: in Lim.ML*) |
|
14329 | 733 |
lemma hypreal_not_eq_minus_iff: "(x \<noteq> a) = (x + -a \<noteq> (0::hypreal))" |
734 |
by (auto dest: hypreal_eq_minus_iff [THEN iffD2]) |
|
735 |
||
736 |
lemma hypreal_mult_left_cancel: "(c::hypreal) \<noteq> 0 ==> (c*a=c*b) = (a=b)" |
|
737 |
apply auto |
|
738 |
done |
|
739 |
||
740 |
lemma hypreal_mult_right_cancel: "(c::hypreal) \<noteq> 0 ==> (a*c=b*c) = (a=b)" |
|
741 |
apply auto |
|
742 |
done |
|
743 |
||
744 |
lemma hypreal_inverse_not_zero: "x \<noteq> 0 ==> inverse (x::hypreal) \<noteq> 0" |
|
745 |
by (rule Ring_and_Field.nonzero_imp_inverse_nonzero) |
|
746 |
||
747 |
lemma hypreal_mult_not_0: "[| x \<noteq> 0; y \<noteq> 0 |] ==> x * y \<noteq> (0::hypreal)" |
|
748 |
by simp |
|
749 |
||
750 |
lemma hypreal_minus_inverse: "inverse(-x) = -inverse(x::hypreal)" |
|
751 |
by (rule Ring_and_Field.inverse_minus_eq) |
|
752 |
||
753 |
lemma hypreal_inverse_distrib: "inverse(x*y) = inverse(x)*inverse(y::hypreal)" |
|
754 |
by (rule Ring_and_Field.inverse_mult_distrib) |
|
755 |
||
756 |
||
757 |
subsection{* Division lemmas *} |
|
758 |
||
759 |
lemma hypreal_divide_one: "x/(1::hypreal) = x" |
|
760 |
by (simp add: hypreal_divide_def) |
|
761 |
||
762 |
||
763 |
(** As with multiplication, pull minus signs OUT of the / operator **) |
|
764 |
||
765 |
lemma hypreal_add_divide_distrib: "(x+y)/(z::hypreal) = x/z + y/z" |
|
766 |
by (rule Ring_and_Field.add_divide_distrib) |
|
767 |
||
768 |
lemma hypreal_inverse_add: |
|
769 |
"[|(x::hypreal) \<noteq> 0; y \<noteq> 0 |] |
|
770 |
==> inverse(x) + inverse(y) = (x + y)*inverse(x*y)" |
|
771 |
by (simp add: Ring_and_Field.inverse_add mult_assoc) |
|
772 |
||
773 |
||
774 |
subsection{*@{term hypreal_of_real} Preserves Field and Order Properties*} |
|
775 |
||
14301 | 776 |
lemma hypreal_of_real_add [simp]: |
14299 | 777 |
"hypreal_of_real (z1 + z2) = hypreal_of_real z1 + hypreal_of_real z2" |
778 |
apply (unfold hypreal_of_real_def) |
|
14331 | 779 |
apply (simp add: hypreal_add left_distrib) |
14299 | 780 |
done |
781 |
||
14301 | 782 |
lemma hypreal_of_real_mult [simp]: |
14299 | 783 |
"hypreal_of_real (z1 * z2) = hypreal_of_real z1 * hypreal_of_real z2" |
784 |
apply (unfold hypreal_of_real_def) |
|
14331 | 785 |
apply (simp add: hypreal_mult right_distrib) |
14299 | 786 |
done |
787 |
||
14301 | 788 |
lemma hypreal_of_real_less_iff [simp]: |
14299 | 789 |
"(hypreal_of_real z1 < hypreal_of_real z2) = (z1 < z2)" |
14301 | 790 |
apply (unfold hypreal_less_def hypreal_of_real_def, auto) |
791 |
apply (rule_tac [2] x = "%n. z1" in exI, safe) |
|
792 |
apply (rule_tac [3] x = "%n. z2" in exI, auto) |
|
793 |
apply (rule FreeUltrafilterNat_P, ultra) |
|
14299 | 794 |
done |
795 |
||
14301 | 796 |
lemma hypreal_of_real_le_iff [simp]: |
14299 | 797 |
"(hypreal_of_real z1 <= hypreal_of_real z2) = (z1 <= z2)" |
14301 | 798 |
apply (unfold hypreal_le_def real_le_def, auto) |
14299 | 799 |
done |
800 |
||
14329 | 801 |
lemma hypreal_of_real_eq_iff [simp]: |
802 |
"(hypreal_of_real z1 = hypreal_of_real z2) = (z1 = z2)" |
|
14301 | 803 |
by (force intro: order_antisym hypreal_of_real_le_iff [THEN iffD1]) |
14299 | 804 |
|
14329 | 805 |
lemma hypreal_of_real_minus [simp]: |
806 |
"hypreal_of_real (-r) = - hypreal_of_real r" |
|
14299 | 807 |
apply (unfold hypreal_of_real_def) |
808 |
apply (auto simp add: hypreal_minus) |
|
809 |
done |
|
810 |
||
14301 | 811 |
lemma hypreal_of_real_one [simp]: "hypreal_of_real 1 = (1::hypreal)" |
812 |
by (unfold hypreal_of_real_def hypreal_one_def, simp) |
|
14299 | 813 |
|
14301 | 814 |
lemma hypreal_of_real_zero [simp]: "hypreal_of_real 0 = 0" |
815 |
by (unfold hypreal_of_real_def hypreal_zero_def, simp) |
|
14299 | 816 |
|
817 |
lemma hypreal_of_real_zero_iff: "(hypreal_of_real r = 0) = (r = 0)" |
|
14301 | 818 |
by (auto intro: FreeUltrafilterNat_P simp add: hypreal_of_real_def hypreal_zero_def FreeUltrafilterNat_Nat_set) |
14299 | 819 |
|
14329 | 820 |
lemma hypreal_of_real_inverse [simp]: |
821 |
"hypreal_of_real (inverse r) = inverse (hypreal_of_real r)" |
|
14299 | 822 |
apply (case_tac "r=0") |
14301 | 823 |
apply (simp add: DIVISION_BY_ZERO INVERSE_ZERO HYPREAL_INVERSE_ZERO) |
14299 | 824 |
apply (rule_tac c1 = "hypreal_of_real r" in hypreal_mult_left_cancel [THEN iffD1]) |
825 |
apply (auto simp add: hypreal_of_real_zero_iff hypreal_of_real_mult [symmetric]) |
|
826 |
done |
|
827 |
||
14329 | 828 |
lemma hypreal_of_real_divide [simp]: |
829 |
"hypreal_of_real (z1 / z2) = hypreal_of_real z1 / hypreal_of_real z2" |
|
14301 | 830 |
by (simp add: hypreal_divide_def real_divide_def) |
14299 | 831 |
|
832 |
||
14329 | 833 |
subsection{*Misc Others*} |
14299 | 834 |
|
835 |
lemma hypreal_zero_num: "0 = Abs_hypreal (hyprel `` {%n. 0})" |
|
14301 | 836 |
by (simp add: hypreal_zero_def [THEN meta_eq_to_obj_eq, symmetric]) |
14299 | 837 |
|
838 |
lemma hypreal_one_num: "1 = Abs_hypreal (hyprel `` {%n. 1})" |
|
14301 | 839 |
by (simp add: hypreal_one_def [THEN meta_eq_to_obj_eq, symmetric]) |
14299 | 840 |
|
14301 | 841 |
lemma hypreal_omega_gt_zero [simp]: "0 < omega" |
14299 | 842 |
apply (unfold omega_def) |
843 |
apply (auto simp add: hypreal_less hypreal_zero_num) |
|
844 |
done |
|
845 |
||
14329 | 846 |
|
847 |
lemma hypreal_hrabs: |
|
848 |
"abs (Abs_hypreal (hyprel `` {X})) = |
|
849 |
Abs_hypreal(hyprel `` {%n. abs (X n)})" |
|
850 |
apply (auto simp add: hrabs_def hypreal_zero_def hypreal_le hypreal_minus) |
|
851 |
apply (ultra, arith)+ |
|
852 |
done |
|
853 |
||
14299 | 854 |
ML |
855 |
{* |
|
14329 | 856 |
val hrabs_def = thm "hrabs_def"; |
857 |
val hypreal_hrabs = thm "hypreal_hrabs"; |
|
858 |
||
14299 | 859 |
val hypreal_zero_def = thm "hypreal_zero_def"; |
860 |
val hypreal_one_def = thm "hypreal_one_def"; |
|
861 |
val hypreal_minus_def = thm "hypreal_minus_def"; |
|
862 |
val hypreal_diff_def = thm "hypreal_diff_def"; |
|
863 |
val hypreal_inverse_def = thm "hypreal_inverse_def"; |
|
864 |
val hypreal_divide_def = thm "hypreal_divide_def"; |
|
865 |
val hypreal_of_real_def = thm "hypreal_of_real_def"; |
|
866 |
val omega_def = thm "omega_def"; |
|
867 |
val epsilon_def = thm "epsilon_def"; |
|
868 |
val hypreal_add_def = thm "hypreal_add_def"; |
|
869 |
val hypreal_mult_def = thm "hypreal_mult_def"; |
|
870 |
val hypreal_less_def = thm "hypreal_less_def"; |
|
871 |
val hypreal_le_def = thm "hypreal_le_def"; |
|
872 |
||
873 |
val finite_exhausts = thm "finite_exhausts"; |
|
874 |
val finite_not_covers = thm "finite_not_covers"; |
|
875 |
val not_finite_nat = thm "not_finite_nat"; |
|
876 |
val FreeUltrafilterNat_Ex = thm "FreeUltrafilterNat_Ex"; |
|
877 |
val FreeUltrafilterNat_mem = thm "FreeUltrafilterNat_mem"; |
|
878 |
val FreeUltrafilterNat_finite = thm "FreeUltrafilterNat_finite"; |
|
879 |
val FreeUltrafilterNat_not_finite = thm "FreeUltrafilterNat_not_finite"; |
|
880 |
val FreeUltrafilterNat_empty = thm "FreeUltrafilterNat_empty"; |
|
881 |
val FreeUltrafilterNat_Int = thm "FreeUltrafilterNat_Int"; |
|
882 |
val FreeUltrafilterNat_subset = thm "FreeUltrafilterNat_subset"; |
|
883 |
val FreeUltrafilterNat_Compl = thm "FreeUltrafilterNat_Compl"; |
|
884 |
val FreeUltrafilterNat_Compl_mem = thm "FreeUltrafilterNat_Compl_mem"; |
|
885 |
val FreeUltrafilterNat_Compl_iff1 = thm "FreeUltrafilterNat_Compl_iff1"; |
|
886 |
val FreeUltrafilterNat_Compl_iff2 = thm "FreeUltrafilterNat_Compl_iff2"; |
|
887 |
val FreeUltrafilterNat_UNIV = thm "FreeUltrafilterNat_UNIV"; |
|
888 |
val FreeUltrafilterNat_Nat_set = thm "FreeUltrafilterNat_Nat_set"; |
|
889 |
val FreeUltrafilterNat_Nat_set_refl = thm "FreeUltrafilterNat_Nat_set_refl"; |
|
890 |
val FreeUltrafilterNat_P = thm "FreeUltrafilterNat_P"; |
|
891 |
val FreeUltrafilterNat_Ex_P = thm "FreeUltrafilterNat_Ex_P"; |
|
892 |
val FreeUltrafilterNat_all = thm "FreeUltrafilterNat_all"; |
|
893 |
val FreeUltrafilterNat_Un = thm "FreeUltrafilterNat_Un"; |
|
894 |
val hyprel_iff = thm "hyprel_iff"; |
|
895 |
val hyprel_refl = thm "hyprel_refl"; |
|
896 |
val hyprel_sym = thm "hyprel_sym"; |
|
897 |
val hyprel_trans = thm "hyprel_trans"; |
|
898 |
val equiv_hyprel = thm "equiv_hyprel"; |
|
899 |
val hyprel_in_hypreal = thm "hyprel_in_hypreal"; |
|
900 |
val Abs_hypreal_inverse = thm "Abs_hypreal_inverse"; |
|
901 |
val inj_on_Abs_hypreal = thm "inj_on_Abs_hypreal"; |
|
902 |
val inj_Rep_hypreal = thm "inj_Rep_hypreal"; |
|
903 |
val lemma_hyprel_refl = thm "lemma_hyprel_refl"; |
|
904 |
val hypreal_empty_not_mem = thm "hypreal_empty_not_mem"; |
|
905 |
val Rep_hypreal_nonempty = thm "Rep_hypreal_nonempty"; |
|
906 |
val inj_hypreal_of_real = thm "inj_hypreal_of_real"; |
|
907 |
val eq_Abs_hypreal = thm "eq_Abs_hypreal"; |
|
908 |
val hypreal_minus_congruent = thm "hypreal_minus_congruent"; |
|
909 |
val hypreal_minus = thm "hypreal_minus"; |
|
910 |
val hypreal_add_congruent2 = thm "hypreal_add_congruent2"; |
|
911 |
val hypreal_add = thm "hypreal_add"; |
|
912 |
val hypreal_diff = thm "hypreal_diff"; |
|
913 |
val hypreal_add_commute = thm "hypreal_add_commute"; |
|
914 |
val hypreal_add_assoc = thm "hypreal_add_assoc"; |
|
915 |
val hypreal_add_zero_left = thm "hypreal_add_zero_left"; |
|
916 |
val hypreal_add_zero_right = thm "hypreal_add_zero_right"; |
|
917 |
val hypreal_add_minus = thm "hypreal_add_minus"; |
|
918 |
val hypreal_add_minus_left = thm "hypreal_add_minus_left"; |
|
919 |
val hypreal_minus_distrib1 = thm "hypreal_minus_distrib1"; |
|
920 |
val hypreal_mult_congruent2 = thm "hypreal_mult_congruent2"; |
|
921 |
val hypreal_mult = thm "hypreal_mult"; |
|
922 |
val hypreal_mult_commute = thm "hypreal_mult_commute"; |
|
923 |
val hypreal_mult_assoc = thm "hypreal_mult_assoc"; |
|
924 |
val hypreal_mult_1 = thm "hypreal_mult_1"; |
|
925 |
val hypreal_mult_1_right = thm "hypreal_mult_1_right"; |
|
926 |
val hypreal_mult_minus_1 = thm "hypreal_mult_minus_1"; |
|
927 |
val hypreal_mult_minus_1_right = thm "hypreal_mult_minus_1_right"; |
|
928 |
val hypreal_zero_not_eq_one = thm "hypreal_zero_not_eq_one"; |
|
929 |
val hypreal_inverse_congruent = thm "hypreal_inverse_congruent"; |
|
930 |
val hypreal_inverse = thm "hypreal_inverse"; |
|
931 |
val HYPREAL_INVERSE_ZERO = thm "HYPREAL_INVERSE_ZERO"; |
|
932 |
val HYPREAL_DIVISION_BY_ZERO = thm "HYPREAL_DIVISION_BY_ZERO"; |
|
933 |
val hypreal_mult_inverse = thm "hypreal_mult_inverse"; |
|
934 |
val hypreal_mult_inverse_left = thm "hypreal_mult_inverse_left"; |
|
935 |
val hypreal_mult_left_cancel = thm "hypreal_mult_left_cancel"; |
|
936 |
val hypreal_mult_right_cancel = thm "hypreal_mult_right_cancel"; |
|
937 |
val hypreal_inverse_not_zero = thm "hypreal_inverse_not_zero"; |
|
938 |
val hypreal_mult_not_0 = thm "hypreal_mult_not_0"; |
|
939 |
val hypreal_minus_inverse = thm "hypreal_minus_inverse"; |
|
940 |
val hypreal_inverse_distrib = thm "hypreal_inverse_distrib"; |
|
941 |
val hypreal_less_not_refl = thm "hypreal_less_not_refl"; |
|
14331 | 942 |
val hypreal_less_irrefl = thm"hypreal_less_irrefl"; |
14299 | 943 |
val hypreal_not_refl2 = thm "hypreal_not_refl2"; |
944 |
val hypreal_less_trans = thm "hypreal_less_trans"; |
|
945 |
val hypreal_less_asym = thm "hypreal_less_asym"; |
|
946 |
val hypreal_less = thm "hypreal_less"; |
|
947 |
val hypreal_trichotomy = thm "hypreal_trichotomy"; |
|
948 |
val hypreal_less_minus_iff = thm "hypreal_less_minus_iff"; |
|
949 |
val hypreal_less_minus_iff2 = thm "hypreal_less_minus_iff2"; |
|
950 |
val hypreal_eq_minus_iff = thm "hypreal_eq_minus_iff"; |
|
951 |
val hypreal_eq_minus_iff2 = thm "hypreal_eq_minus_iff2"; |
|
952 |
val hypreal_eq_minus_iff3 = thm "hypreal_eq_minus_iff3"; |
|
953 |
val hypreal_not_eq_minus_iff = thm "hypreal_not_eq_minus_iff"; |
|
954 |
val hypreal_linear = thm "hypreal_linear"; |
|
955 |
val hypreal_neq_iff = thm "hypreal_neq_iff"; |
|
956 |
val hypreal_linear_less2 = thm "hypreal_linear_less2"; |
|
957 |
val hypreal_le = thm "hypreal_le"; |
|
958 |
val hypreal_le_imp_less_or_eq = thm "hypreal_le_imp_less_or_eq"; |
|
959 |
val hypreal_le_eq_less_or_eq = thm "hypreal_le_eq_less_or_eq"; |
|
960 |
val hypreal_le_refl = thm "hypreal_le_refl"; |
|
961 |
val hypreal_le_linear = thm "hypreal_le_linear"; |
|
962 |
val hypreal_le_trans = thm "hypreal_le_trans"; |
|
963 |
val hypreal_le_anti_sym = thm "hypreal_le_anti_sym"; |
|
964 |
val hypreal_less_le = thm "hypreal_less_le"; |
|
965 |
val hypreal_of_real_add = thm "hypreal_of_real_add"; |
|
966 |
val hypreal_of_real_mult = thm "hypreal_of_real_mult"; |
|
967 |
val hypreal_of_real_less_iff = thm "hypreal_of_real_less_iff"; |
|
968 |
val hypreal_of_real_le_iff = thm "hypreal_of_real_le_iff"; |
|
969 |
val hypreal_of_real_eq_iff = thm "hypreal_of_real_eq_iff"; |
|
970 |
val hypreal_of_real_minus = thm "hypreal_of_real_minus"; |
|
971 |
val hypreal_of_real_one = thm "hypreal_of_real_one"; |
|
972 |
val hypreal_of_real_zero = thm "hypreal_of_real_zero"; |
|
973 |
val hypreal_of_real_zero_iff = thm "hypreal_of_real_zero_iff"; |
|
974 |
val hypreal_of_real_inverse = thm "hypreal_of_real_inverse"; |
|
975 |
val hypreal_of_real_divide = thm "hypreal_of_real_divide"; |
|
976 |
val hypreal_divide_one = thm "hypreal_divide_one"; |
|
977 |
val hypreal_add_divide_distrib = thm "hypreal_add_divide_distrib"; |
|
978 |
val hypreal_inverse_add = thm "hypreal_inverse_add"; |
|
979 |
val hypreal_zero_num = thm "hypreal_zero_num"; |
|
980 |
val hypreal_one_num = thm "hypreal_one_num"; |
|
981 |
val hypreal_omega_gt_zero = thm "hypreal_omega_gt_zero"; |
|
982 |
*} |
|
983 |
||
10751 | 984 |
end |