| author | haftmann |
| Mon, 25 Sep 2006 17:04:17 +0200 | |
| changeset 20700 | 7e3450c10c2d |
| parent 20652 | 6e9b7617c89a |
| child 20713 | 823967ef47f1 |
| permissions | -rw-r--r-- |
| 10751 | 1 |
(* Title : NSA.thy |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 1998 University of Cambridge |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
4 |
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
5 |
Converted to Isar and polished by lcp |
| 14370 | 6 |
*) |
| 10751 | 7 |
|
| 14370 | 8 |
header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
|
9 |
||
| 15131 | 10 |
theory NSA |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
11 |
imports HyperArith "../Real/RComplete" |
| 15131 | 12 |
begin |
| 10751 | 13 |
|
| 19765 | 14 |
definition |
| 10751 | 15 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
16 |
hnorm :: "'a::norm star \<Rightarrow> real star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
17 |
"hnorm = *f* norm" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
18 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
19 |
Infinitesimal :: "('a::real_normed_vector) star set"
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
20 |
"Infinitesimal = {x. \<forall>r \<in> Reals. 0 < r --> hnorm x < r}"
|
| 10751 | 21 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
22 |
HFinite :: "('a::real_normed_vector) star set"
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
23 |
"HFinite = {x. \<exists>r \<in> Reals. hnorm x < r}"
|
| 10751 | 24 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
25 |
HInfinite :: "('a::real_normed_vector) star set"
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
26 |
"HInfinite = {x. \<forall>r \<in> Reals. r < hnorm x}"
|
| 10751 | 27 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
28 |
approx :: "['a::real_normed_vector star, 'a star] => bool" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
29 |
(infixl "@=" 50) |
| 15229 | 30 |
--{*the `infinitely close' relation*}
|
| 20563 | 31 |
"(x @= y) = ((x - y) \<in> Infinitesimal)" |
| 14653 | 32 |
|
| 14370 | 33 |
st :: "hypreal => hypreal" |
| 15229 | 34 |
--{*the standard part of a hyperreal*}
|
| 19765 | 35 |
"st = (%x. @r. x \<in> HFinite & r \<in> Reals & r @= x)" |
| 10751 | 36 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
37 |
monad :: "'a::real_normed_vector star => 'a star set" |
| 19765 | 38 |
"monad x = {y. x @= y}"
|
| 10751 | 39 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
40 |
galaxy :: "'a::real_normed_vector star => 'a star set" |
| 19765 | 41 |
"galaxy x = {y. (x + -y) \<in> HFinite}"
|
42 |
||
43 |
const_syntax (xsymbols) |
|
44 |
approx (infixl "\<approx>" 50) |
|
45 |
||
46 |
const_syntax (HTML output) |
|
47 |
approx (infixl "\<approx>" 50) |
|
| 14370 | 48 |
|
49 |
||
|
20554
c433e78d4203
define new constant of_real for class real_algebra_1;
huffman
parents:
20552
diff
changeset
|
50 |
lemma hypreal_of_real_of_real_eq: "hypreal_of_real r = of_real r" |
|
c433e78d4203
define new constant of_real for class real_algebra_1;
huffman
parents:
20552
diff
changeset
|
51 |
proof - |
|
c433e78d4203
define new constant of_real for class real_algebra_1;
huffman
parents:
20552
diff
changeset
|
52 |
have "hypreal_of_real r = hypreal_of_real (of_real r)" by simp |
|
c433e78d4203
define new constant of_real for class real_algebra_1;
huffman
parents:
20552
diff
changeset
|
53 |
also have "\<dots> = of_real r" by (rule star_of_of_real) |
|
c433e78d4203
define new constant of_real for class real_algebra_1;
huffman
parents:
20552
diff
changeset
|
54 |
finally show ?thesis . |
|
c433e78d4203
define new constant of_real for class real_algebra_1;
huffman
parents:
20552
diff
changeset
|
55 |
qed |
| 14565 | 56 |
|
|
20554
c433e78d4203
define new constant of_real for class real_algebra_1;
huffman
parents:
20552
diff
changeset
|
57 |
lemma SReal_def: "Reals == {x. \<exists>r. x = hypreal_of_real r}"
|
|
c433e78d4203
define new constant of_real for class real_algebra_1;
huffman
parents:
20552
diff
changeset
|
58 |
by (simp add: Reals_def image_def hypreal_of_real_of_real_eq) |
| 14370 | 59 |
|
60 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
61 |
subsection{*Nonstandard extension of the norm function*}
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
62 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
63 |
declare hnorm_def [transfer_unfold] |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
64 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
65 |
lemma hnorm_ge_zero [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
66 |
"\<And>x::'a::real_normed_vector star. 0 \<le> hnorm x" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
67 |
by transfer (rule norm_ge_zero) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
68 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
69 |
lemma hnorm_eq_zero [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
70 |
"\<And>x::'a::real_normed_vector star. (hnorm x = 0) = (x = 0)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
71 |
by transfer (rule norm_eq_zero) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
72 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
73 |
lemma hnorm_triangle_ineq: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
74 |
"\<And>x y::'a::real_normed_vector star. hnorm (x + y) \<le> hnorm x + hnorm y" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
75 |
by transfer (rule norm_triangle_ineq) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
76 |
|
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
77 |
lemma hnorm_triangle_ineq3: |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
78 |
"\<And>x y::'a::real_normed_vector star. \<bar>hnorm x - hnorm y\<bar> \<le> hnorm (x - y)" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
79 |
by transfer (rule norm_triangle_ineq3) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
80 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
81 |
lemma hnorm_scaleR: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
82 |
"\<And>a (x::'a::real_normed_vector star). |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
83 |
hnorm (( *f2* scaleR) a x) = \<bar>a\<bar> * hnorm x" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
84 |
by transfer (rule norm_scaleR) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
85 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
86 |
lemma hnorm_mult_ineq: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
87 |
"\<And>x y::'a::real_normed_algebra star. hnorm (x * y) \<le> hnorm x * hnorm y" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
88 |
by transfer (rule norm_mult_ineq) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
89 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
90 |
lemma hnorm_mult: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
91 |
"\<And>x y::'a::real_normed_div_algebra star. hnorm (x * y) = hnorm x * hnorm y" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
92 |
by transfer (rule norm_mult) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
93 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
94 |
lemma hnorm_one [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
95 |
"hnorm (1\<Colon>'a::real_normed_div_algebra star) = 1" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
96 |
by transfer (rule norm_one) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
97 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
98 |
lemma hnorm_zero [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
99 |
"hnorm (0\<Colon>'a::real_normed_vector star) = 0" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
100 |
by transfer (rule norm_zero) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
101 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
102 |
lemma zero_less_hnorm_iff [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
103 |
"\<And>x::'a::real_normed_vector star. (0 < hnorm x) = (x \<noteq> 0)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
104 |
by transfer (rule zero_less_norm_iff) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
105 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
106 |
lemma hnorm_minus_cancel [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
107 |
"\<And>x::'a::real_normed_vector star. hnorm (- x) = hnorm x" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
108 |
by transfer (rule norm_minus_cancel) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
109 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
110 |
lemma hnorm_minus_commute: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
111 |
"\<And>a b::'a::real_normed_vector star. hnorm (a - b) = hnorm (b - a)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
112 |
by transfer (rule norm_minus_commute) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
113 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
114 |
lemma hnorm_triangle_ineq2: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
115 |
"\<And>a b::'a::real_normed_vector star. hnorm a - hnorm b \<le> hnorm (a - b)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
116 |
by transfer (rule norm_triangle_ineq2) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
117 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
118 |
lemma hnorm_triangle_ineq4: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
119 |
"\<And>a b::'a::real_normed_vector star. hnorm (a - b) \<le> hnorm a + hnorm b" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
120 |
by transfer (rule norm_triangle_ineq4) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
121 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
122 |
lemma nonzero_hnorm_inverse: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
123 |
"\<And>a::'a::real_normed_div_algebra star. |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
124 |
a \<noteq> 0 \<Longrightarrow> hnorm (inverse a) = inverse (hnorm a)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
125 |
by transfer (rule nonzero_norm_inverse) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
126 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
127 |
lemma hnorm_inverse: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
128 |
"\<And>a::'a::{real_normed_div_algebra,division_by_zero} star.
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
129 |
hnorm (inverse a) = inverse (hnorm a)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
130 |
by transfer (rule norm_inverse) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
131 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
132 |
lemma hypreal_hnorm_def [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
133 |
"\<And>r::hypreal. hnorm r \<equiv> \<bar>r\<bar>" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
134 |
by transfer (rule real_norm_def) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
135 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
136 |
lemma hnorm_add_less: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
137 |
fixes x y :: "'a::real_normed_vector star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
138 |
shows "\<lbrakk>hnorm x < r; hnorm y < s\<rbrakk> \<Longrightarrow> hnorm (x + y) < r + s" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
139 |
by (rule order_le_less_trans [OF hnorm_triangle_ineq add_strict_mono]) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
140 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
141 |
lemma hnorm_mult_less: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
142 |
fixes x y :: "'a::real_normed_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
143 |
shows "\<lbrakk>hnorm x < r; hnorm y < s\<rbrakk> \<Longrightarrow> hnorm (x * y) < r * s" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
144 |
apply (rule order_le_less_trans [OF hnorm_mult_ineq]) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
145 |
apply (simp add: mult_strict_mono') |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
146 |
done |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
147 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
148 |
|
| 15229 | 149 |
subsection{*Closure Laws for the Standard Reals*}
|
| 14370 | 150 |
|
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
151 |
lemma SReal_add [simp]: |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
152 |
"[| (x::hypreal) \<in> Reals; y \<in> Reals |] ==> x + y \<in> Reals" |
| 14370 | 153 |
apply (auto simp add: SReal_def) |
154 |
apply (rule_tac x = "r + ra" in exI, simp) |
|
155 |
done |
|
156 |
||
| 20563 | 157 |
lemma SReal_diff [simp]: |
158 |
"[| (x::hypreal) \<in> Reals; y \<in> Reals |] ==> x - y \<in> Reals" |
|
159 |
apply (auto simp add: SReal_def) |
|
160 |
apply (rule_tac x = "r - ra" in exI, simp) |
|
161 |
done |
|
162 |
||
| 14370 | 163 |
lemma SReal_mult: "[| (x::hypreal) \<in> Reals; y \<in> Reals |] ==> x * y \<in> Reals" |
164 |
apply (simp add: SReal_def, safe) |
|
165 |
apply (rule_tac x = "r * ra" in exI) |
|
| 15539 | 166 |
apply (simp (no_asm)) |
| 14370 | 167 |
done |
168 |
||
169 |
lemma SReal_inverse: "(x::hypreal) \<in> Reals ==> inverse x \<in> Reals" |
|
170 |
apply (simp add: SReal_def) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
171 |
apply (blast intro: star_of_inverse [symmetric]) |
| 14370 | 172 |
done |
173 |
||
174 |
lemma SReal_divide: "[| (x::hypreal) \<in> Reals; y \<in> Reals |] ==> x/y \<in> Reals" |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
175 |
by (simp (no_asm_simp) add: SReal_mult SReal_inverse divide_inverse) |
| 14370 | 176 |
|
177 |
lemma SReal_minus: "(x::hypreal) \<in> Reals ==> -x \<in> Reals" |
|
178 |
apply (simp add: SReal_def) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
179 |
apply (blast intro: star_of_minus [symmetric]) |
| 14370 | 180 |
done |
181 |
||
| 15229 | 182 |
lemma SReal_minus_iff [simp]: "(-x \<in> Reals) = ((x::hypreal) \<in> Reals)" |
| 14370 | 183 |
apply auto |
184 |
apply (drule SReal_minus, auto) |
|
185 |
done |
|
186 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
187 |
lemma SReal_add_cancel: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
188 |
"[| (x::hypreal) + y \<in> Reals; y \<in> Reals |] ==> x \<in> Reals" |
| 14370 | 189 |
apply (drule_tac x = y in SReal_minus) |
190 |
apply (drule SReal_add, assumption, auto) |
|
191 |
done |
|
192 |
||
193 |
lemma SReal_hrabs: "(x::hypreal) \<in> Reals ==> abs x \<in> Reals" |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
194 |
apply (auto simp add: SReal_def) |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
195 |
apply (rule_tac x="abs r" in exI) |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
196 |
apply simp |
| 14370 | 197 |
done |
198 |
||
| 15229 | 199 |
lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \<in> Reals" |
| 14370 | 200 |
by (simp add: SReal_def) |
201 |
||
| 15229 | 202 |
lemma SReal_number_of [simp]: "(number_of w ::hypreal) \<in> Reals" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
203 |
apply (simp only: star_of_number_of [symmetric]) |
| 14370 | 204 |
apply (rule SReal_hypreal_of_real) |
205 |
done |
|
206 |
||
207 |
(** As always with numerals, 0 and 1 are special cases **) |
|
208 |
||
| 15229 | 209 |
lemma Reals_0 [simp]: "(0::hypreal) \<in> Reals" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
210 |
apply (subst numeral_0_eq_0 [symmetric]) |
| 14370 | 211 |
apply (rule SReal_number_of) |
212 |
done |
|
213 |
||
| 15229 | 214 |
lemma Reals_1 [simp]: "(1::hypreal) \<in> Reals" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
215 |
apply (subst numeral_1_eq_1 [symmetric]) |
| 14370 | 216 |
apply (rule SReal_number_of) |
217 |
done |
|
218 |
||
219 |
lemma SReal_divide_number_of: "r \<in> Reals ==> r/(number_of w::hypreal) \<in> Reals" |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
220 |
apply (simp only: divide_inverse) |
| 14370 | 221 |
apply (blast intro!: SReal_number_of SReal_mult SReal_inverse) |
222 |
done |
|
223 |
||
| 15229 | 224 |
text{*epsilon is not in Reals because it is an infinitesimal*}
|
| 14370 | 225 |
lemma SReal_epsilon_not_mem: "epsilon \<notin> Reals" |
226 |
apply (simp add: SReal_def) |
|
227 |
apply (auto simp add: hypreal_of_real_not_eq_epsilon [THEN not_sym]) |
|
228 |
done |
|
229 |
||
230 |
lemma SReal_omega_not_mem: "omega \<notin> Reals" |
|
231 |
apply (simp add: SReal_def) |
|
232 |
apply (auto simp add: hypreal_of_real_not_eq_omega [THEN not_sym]) |
|
233 |
done |
|
234 |
||
235 |
lemma SReal_UNIV_real: "{x. hypreal_of_real x \<in> Reals} = (UNIV::real set)"
|
|
236 |
by (simp add: SReal_def) |
|
237 |
||
238 |
lemma SReal_iff: "(x \<in> Reals) = (\<exists>y. x = hypreal_of_real y)" |
|
239 |
by (simp add: SReal_def) |
|
240 |
||
241 |
lemma hypreal_of_real_image: "hypreal_of_real `(UNIV::real set) = Reals" |
|
242 |
by (auto simp add: SReal_def) |
|
243 |
||
244 |
lemma inv_hypreal_of_real_image: "inv hypreal_of_real ` Reals = UNIV" |
|
245 |
apply (auto simp add: SReal_def) |
|
246 |
apply (rule inj_hypreal_of_real [THEN inv_f_f, THEN subst], blast) |
|
247 |
done |
|
248 |
||
249 |
lemma SReal_hypreal_of_real_image: |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
250 |
"[| \<exists>x. x: P; P \<subseteq> Reals |] ==> \<exists>Q. P = hypreal_of_real ` Q" |
| 14370 | 251 |
apply (simp add: SReal_def, blast) |
252 |
done |
|
253 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
254 |
lemma SReal_dense: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
255 |
"[| (x::hypreal) \<in> Reals; y \<in> Reals; x<y |] ==> \<exists>r \<in> Reals. x<r & r<y" |
| 14370 | 256 |
apply (auto simp add: SReal_iff) |
| 14477 | 257 |
apply (drule dense, safe) |
| 14370 | 258 |
apply (rule_tac x = "hypreal_of_real r" in bexI, auto) |
259 |
done |
|
260 |
||
| 15229 | 261 |
text{*Completeness of Reals, but both lemmas are unused.*}
|
262 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
263 |
lemma SReal_sup_lemma: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
264 |
"P \<subseteq> Reals ==> ((\<exists>x \<in> P. y < x) = |
| 14370 | 265 |
(\<exists>X. hypreal_of_real X \<in> P & y < hypreal_of_real X))" |
266 |
by (blast dest!: SReal_iff [THEN iffD1]) |
|
267 |
||
268 |
lemma SReal_sup_lemma2: |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
269 |
"[| P \<subseteq> Reals; \<exists>x. x \<in> P; \<exists>y \<in> Reals. \<forall>x \<in> P. x < y |] |
| 14370 | 270 |
==> (\<exists>X. X \<in> {w. hypreal_of_real w \<in> P}) &
|
271 |
(\<exists>Y. \<forall>X \<in> {w. hypreal_of_real w \<in> P}. X < Y)"
|
|
272 |
apply (rule conjI) |
|
273 |
apply (fast dest!: SReal_iff [THEN iffD1]) |
|
274 |
apply (auto, frule subsetD, assumption) |
|
275 |
apply (drule SReal_iff [THEN iffD1]) |
|
276 |
apply (auto, rule_tac x = ya in exI, auto) |
|
277 |
done |
|
278 |
||
| 15229 | 279 |
|
280 |
subsection{*Lifting of the Ub and Lub Properties*}
|
|
281 |
||
| 14370 | 282 |
lemma hypreal_of_real_isUb_iff: |
283 |
"(isUb (Reals) (hypreal_of_real ` Q) (hypreal_of_real Y)) = |
|
284 |
(isUb (UNIV :: real set) Q Y)" |
|
| 15229 | 285 |
by (simp add: isUb_def setle_def) |
| 14370 | 286 |
|
287 |
lemma hypreal_of_real_isLub1: |
|
288 |
"isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y) |
|
289 |
==> isLub (UNIV :: real set) Q Y" |
|
290 |
apply (simp add: isLub_def leastP_def) |
|
291 |
apply (auto intro: hypreal_of_real_isUb_iff [THEN iffD2] |
|
292 |
simp add: hypreal_of_real_isUb_iff setge_def) |
|
293 |
done |
|
294 |
||
295 |
lemma hypreal_of_real_isLub2: |
|
296 |
"isLub (UNIV :: real set) Q Y |
|
297 |
==> isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)" |
|
298 |
apply (simp add: isLub_def leastP_def) |
|
299 |
apply (auto simp add: hypreal_of_real_isUb_iff setge_def) |
|
300 |
apply (frule_tac x2 = x in isUbD2a [THEN SReal_iff [THEN iffD1], THEN exE]) |
|
301 |
prefer 2 apply assumption |
|
302 |
apply (drule_tac x = xa in spec) |
|
303 |
apply (auto simp add: hypreal_of_real_isUb_iff) |
|
304 |
done |
|
305 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
306 |
lemma hypreal_of_real_isLub_iff: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
307 |
"(isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)) = |
| 14370 | 308 |
(isLub (UNIV :: real set) Q Y)" |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
309 |
by (blast intro: hypreal_of_real_isLub1 hypreal_of_real_isLub2) |
| 14370 | 310 |
|
311 |
lemma lemma_isUb_hypreal_of_real: |
|
312 |
"isUb Reals P Y ==> \<exists>Yo. isUb Reals P (hypreal_of_real Yo)" |
|
313 |
by (auto simp add: SReal_iff isUb_def) |
|
314 |
||
315 |
lemma lemma_isLub_hypreal_of_real: |
|
316 |
"isLub Reals P Y ==> \<exists>Yo. isLub Reals P (hypreal_of_real Yo)" |
|
317 |
by (auto simp add: isLub_def leastP_def isUb_def SReal_iff) |
|
318 |
||
319 |
lemma lemma_isLub_hypreal_of_real2: |
|
320 |
"\<exists>Yo. isLub Reals P (hypreal_of_real Yo) ==> \<exists>Y. isLub Reals P Y" |
|
321 |
by (auto simp add: isLub_def leastP_def isUb_def) |
|
322 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
323 |
lemma SReal_complete: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
324 |
"[| P \<subseteq> Reals; \<exists>x. x \<in> P; \<exists>Y. isUb Reals P Y |] |
| 14370 | 325 |
==> \<exists>t::hypreal. isLub Reals P t" |
326 |
apply (frule SReal_hypreal_of_real_image) |
|
327 |
apply (auto, drule lemma_isUb_hypreal_of_real) |
|
| 15229 | 328 |
apply (auto intro!: reals_complete lemma_isLub_hypreal_of_real2 |
329 |
simp add: hypreal_of_real_isLub_iff hypreal_of_real_isUb_iff) |
|
| 14370 | 330 |
done |
331 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
332 |
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
333 |
subsection{* Set of Finite Elements is a Subring of the Extended Reals*}
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
334 |
|
| 14370 | 335 |
lemma HFinite_add: "[|x \<in> HFinite; y \<in> HFinite|] ==> (x+y) \<in> HFinite" |
336 |
apply (simp add: HFinite_def) |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
337 |
apply (blast intro!: SReal_add hnorm_add_less) |
| 14370 | 338 |
done |
339 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
340 |
lemma HFinite_mult: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
341 |
fixes x y :: "'a::real_normed_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
342 |
shows "[|x \<in> HFinite; y \<in> HFinite|] ==> x*y \<in> HFinite" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
343 |
apply (simp add: HFinite_def) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
344 |
apply (blast intro!: SReal_mult hnorm_mult_less) |
| 14370 | 345 |
done |
346 |
||
347 |
lemma HFinite_minus_iff: "(-x \<in> HFinite) = (x \<in> HFinite)" |
|
348 |
by (simp add: HFinite_def) |
|
349 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
350 |
lemma HFinite_star_of [simp]: "star_of x \<in> HFinite" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
351 |
apply (simp add: HFinite_def) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
352 |
apply (rule_tac x="star_of (norm x) + 1" in bexI) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
353 |
apply (transfer, simp) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
354 |
apply (blast intro: SReal_add SReal_hypreal_of_real Reals_1) |
| 14370 | 355 |
done |
356 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
357 |
lemma SReal_subset_HFinite: "(Reals::hypreal set) \<subseteq> HFinite" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
358 |
by (auto simp add: SReal_def) |
| 14370 | 359 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
360 |
lemma HFinite_hypreal_of_real: "hypreal_of_real x \<in> HFinite" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
361 |
by (rule HFinite_star_of) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
362 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
363 |
lemma HFiniteD: "x \<in> HFinite ==> \<exists>t \<in> Reals. hnorm x < t" |
| 14370 | 364 |
by (simp add: HFinite_def) |
365 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
366 |
lemma HFinite_hrabs_iff [iff]: "(abs (x::hypreal) \<in> HFinite) = (x \<in> HFinite)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
367 |
by (simp add: HFinite_def) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
368 |
|
|
20562
c2f672be8522
add type constraint to otherwise looping iff rule
huffman
parents:
20554
diff
changeset
|
369 |
lemma HFinite_hnorm_iff [iff]: |
|
c2f672be8522
add type constraint to otherwise looping iff rule
huffman
parents:
20554
diff
changeset
|
370 |
"(hnorm (x::hypreal) \<in> HFinite) = (x \<in> HFinite)" |
| 14370 | 371 |
by (simp add: HFinite_def) |
372 |
||
| 15229 | 373 |
lemma HFinite_number_of [simp]: "number_of w \<in> HFinite" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
374 |
by (unfold star_number_def, rule HFinite_star_of) |
| 14370 | 375 |
|
376 |
(** As always with numerals, 0 and 1 are special cases **) |
|
377 |
||
| 15229 | 378 |
lemma HFinite_0 [simp]: "0 \<in> HFinite" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
379 |
by (unfold star_zero_def, rule HFinite_star_of) |
| 14370 | 380 |
|
| 15229 | 381 |
lemma HFinite_1 [simp]: "1 \<in> HFinite" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
382 |
by (unfold star_one_def, rule HFinite_star_of) |
| 14370 | 383 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
384 |
lemma HFinite_bounded: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
385 |
"[|(x::hypreal) \<in> HFinite; y \<le> x; 0 \<le> y |] ==> y \<in> HFinite" |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
386 |
apply (case_tac "x \<le> 0") |
| 14370 | 387 |
apply (drule_tac y = x in order_trans) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
388 |
apply (drule_tac [2] order_antisym) |
| 14370 | 389 |
apply (auto simp add: linorder_not_le) |
390 |
apply (auto intro: order_le_less_trans simp add: abs_if HFinite_def) |
|
391 |
done |
|
392 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
393 |
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
394 |
subsection{* Set of Infinitesimals is a Subring of the Hyperreals*}
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
395 |
|
| 20407 | 396 |
lemma InfinitesimalI: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
397 |
"(\<And>r. \<lbrakk>r \<in> \<real>; 0 < r\<rbrakk> \<Longrightarrow> hnorm x < r) \<Longrightarrow> x \<in> Infinitesimal" |
| 20407 | 398 |
by (simp add: Infinitesimal_def) |
399 |
||
| 14370 | 400 |
lemma InfinitesimalD: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
401 |
"x \<in> Infinitesimal ==> \<forall>r \<in> Reals. 0 < r --> hnorm x < r" |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
402 |
by (simp add: Infinitesimal_def) |
| 14370 | 403 |
|
| 15229 | 404 |
lemma Infinitesimal_zero [iff]: "0 \<in> Infinitesimal" |
| 14370 | 405 |
by (simp add: Infinitesimal_def) |
406 |
||
407 |
lemma hypreal_sum_of_halves: "x/(2::hypreal) + x/(2::hypreal) = x" |
|
408 |
by auto |
|
409 |
||
410 |
lemma Infinitesimal_add: |
|
411 |
"[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> (x+y) \<in> Infinitesimal" |
|
| 20407 | 412 |
apply (rule InfinitesimalI) |
| 14370 | 413 |
apply (rule hypreal_sum_of_halves [THEN subst]) |
| 14477 | 414 |
apply (drule half_gt_zero) |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
415 |
apply (blast intro: hnorm_add_less SReal_divide_number_of dest: InfinitesimalD) |
| 14370 | 416 |
done |
417 |
||
| 15229 | 418 |
lemma Infinitesimal_minus_iff [simp]: "(-x:Infinitesimal) = (x:Infinitesimal)" |
| 14370 | 419 |
by (simp add: Infinitesimal_def) |
420 |
||
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
421 |
lemma Infinitesimal_hnorm_iff: |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
422 |
"(hnorm x \<in> Infinitesimal) = (x \<in> Infinitesimal)" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
423 |
by (simp add: Infinitesimal_def) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
424 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
425 |
lemma Infinitesimal_diff: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
426 |
"[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> x-y \<in> Infinitesimal" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
427 |
by (simp add: diff_def Infinitesimal_add) |
| 14370 | 428 |
|
429 |
lemma Infinitesimal_mult: |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
430 |
fixes x y :: "'a::real_normed_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
431 |
shows "[|x \<in> Infinitesimal; y \<in> Infinitesimal|] ==> (x * y) \<in> Infinitesimal" |
| 20407 | 432 |
apply (rule InfinitesimalI) |
433 |
apply (case_tac "y = 0", simp) |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
434 |
apply (subgoal_tac "hnorm (x * y) < 1 * r", simp only: mult_1) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
435 |
apply (rule hnorm_mult_less) |
| 20407 | 436 |
apply (simp_all add: InfinitesimalD) |
| 14370 | 437 |
done |
438 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
439 |
lemma Infinitesimal_HFinite_mult: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
440 |
fixes x y :: "'a::real_normed_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
441 |
shows "[| x \<in> Infinitesimal; y \<in> HFinite |] ==> (x * y) \<in> Infinitesimal" |
| 20407 | 442 |
apply (rule InfinitesimalI) |
443 |
apply (drule HFiniteD, clarify) |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
444 |
apply (subgoal_tac "0 < t") |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
445 |
apply (subgoal_tac "hnorm (x * y) < (r / t) * t", simp) |
| 20407 | 446 |
apply (subgoal_tac "0 < r / t") |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
447 |
apply (rule hnorm_mult_less) |
| 20407 | 448 |
apply (simp add: InfinitesimalD SReal_divide) |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
449 |
apply assumption |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
450 |
apply (simp only: divide_pos_pos) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
451 |
apply (erule order_le_less_trans [OF hnorm_ge_zero]) |
| 14370 | 452 |
done |
453 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
454 |
lemma Infinitesimal_HFinite_mult2: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
455 |
fixes x y :: "'a::real_normed_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
456 |
shows "[| x \<in> Infinitesimal; y \<in> HFinite |] ==> (y * x) \<in> Infinitesimal" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
457 |
apply (rule InfinitesimalI) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
458 |
apply (drule HFiniteD, clarify) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
459 |
apply (subgoal_tac "0 < t") |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
460 |
apply (subgoal_tac "hnorm (y * x) < t * (r / t)", simp) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
461 |
apply (subgoal_tac "0 < r / t") |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
462 |
apply (rule hnorm_mult_less) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
463 |
apply assumption |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
464 |
apply (simp add: InfinitesimalD SReal_divide) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
465 |
apply (simp only: divide_pos_pos) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
466 |
apply (erule order_le_less_trans [OF hnorm_ge_zero]) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
467 |
done |
| 14370 | 468 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
469 |
lemma Compl_HFinite: "- HFinite = HInfinite" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
470 |
apply (auto simp add: HInfinite_def HFinite_def linorder_not_less) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
471 |
apply (rule_tac y="r + 1" in order_less_le_trans, simp) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
472 |
apply (simp add: SReal_add Reals_1) |
| 14370 | 473 |
done |
474 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
475 |
lemma HInfinite_inverse_Infinitesimal: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
476 |
fixes x :: "'a::real_normed_div_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
477 |
shows "x \<in> HInfinite ==> inverse x \<in> Infinitesimal" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
478 |
apply (rule InfinitesimalI) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
479 |
apply (subgoal_tac "x \<noteq> 0") |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
480 |
apply (rule inverse_less_imp_less) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
481 |
apply (simp add: nonzero_hnorm_inverse) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
482 |
apply (simp add: HInfinite_def SReal_inverse) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
483 |
apply assumption |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
484 |
apply (clarify, simp add: Compl_HFinite [symmetric]) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
485 |
done |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
486 |
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
487 |
lemma HInfiniteI: "(\<And>r. r \<in> \<real> \<Longrightarrow> r < hnorm x) \<Longrightarrow> x \<in> HInfinite" |
| 20407 | 488 |
by (simp add: HInfinite_def) |
489 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
490 |
lemma HInfiniteD: "\<lbrakk>x \<in> HInfinite; r \<in> \<real>\<rbrakk> \<Longrightarrow> r < hnorm x" |
| 20407 | 491 |
by (simp add: HInfinite_def) |
492 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
493 |
lemma HInfinite_mult: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
494 |
fixes x y :: "'a::real_normed_div_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
495 |
shows "[|x \<in> HInfinite; y \<in> HInfinite|] ==> (x*y) \<in> HInfinite" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
496 |
apply (rule HInfiniteI, simp only: hnorm_mult) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
497 |
apply (subgoal_tac "r * 1 < hnorm x * hnorm y", simp only: mult_1) |
| 20407 | 498 |
apply (case_tac "x = 0", simp add: HInfinite_def) |
499 |
apply (rule mult_strict_mono) |
|
500 |
apply (simp_all add: HInfiniteD) |
|
| 14370 | 501 |
done |
502 |
||
| 15229 | 503 |
lemma hypreal_add_zero_less_le_mono: "[|r < x; (0::hypreal) \<le> y|] ==> r < x+y" |
504 |
by (auto dest: add_less_le_mono) |
|
505 |
||
| 14370 | 506 |
lemma HInfinite_add_ge_zero: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
507 |
"[|(x::hypreal) \<in> HInfinite; 0 \<le> y; 0 \<le> x|] ==> (x + y): HInfinite" |
| 14370 | 508 |
by (auto intro!: hypreal_add_zero_less_le_mono |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
509 |
simp add: abs_if add_commute add_nonneg_nonneg HInfinite_def) |
| 14370 | 510 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
511 |
lemma HInfinite_add_ge_zero2: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
512 |
"[|(x::hypreal) \<in> HInfinite; 0 \<le> y; 0 \<le> x|] ==> (y + x): HInfinite" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
513 |
by (auto intro!: HInfinite_add_ge_zero simp add: add_commute) |
| 14370 | 514 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
515 |
lemma HInfinite_add_gt_zero: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
516 |
"[|(x::hypreal) \<in> HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite" |
| 14370 | 517 |
by (blast intro: HInfinite_add_ge_zero order_less_imp_le) |
518 |
||
519 |
lemma HInfinite_minus_iff: "(-x \<in> HInfinite) = (x \<in> HInfinite)" |
|
520 |
by (simp add: HInfinite_def) |
|
521 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
522 |
lemma HInfinite_add_le_zero: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
523 |
"[|(x::hypreal) \<in> HInfinite; y \<le> 0; x \<le> 0|] ==> (x + y): HInfinite" |
| 14370 | 524 |
apply (drule HInfinite_minus_iff [THEN iffD2]) |
525 |
apply (rule HInfinite_minus_iff [THEN iffD1]) |
|
526 |
apply (auto intro: HInfinite_add_ge_zero) |
|
527 |
done |
|
528 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
529 |
lemma HInfinite_add_lt_zero: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
530 |
"[|(x::hypreal) \<in> HInfinite; y < 0; x < 0|] ==> (x + y): HInfinite" |
| 14370 | 531 |
by (blast intro: HInfinite_add_le_zero order_less_imp_le) |
532 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
533 |
lemma HFinite_sum_squares: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
534 |
fixes a b c :: "'a::real_normed_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
535 |
shows "[|a: HFinite; b: HFinite; c: HFinite|] |
| 14370 | 536 |
==> a*a + b*b + c*c \<in> HFinite" |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
537 |
by (auto intro: HFinite_mult HFinite_add) |
| 14370 | 538 |
|
539 |
lemma not_Infinitesimal_not_zero: "x \<notin> Infinitesimal ==> x \<noteq> 0" |
|
540 |
by auto |
|
541 |
||
542 |
lemma not_Infinitesimal_not_zero2: "x \<in> HFinite - Infinitesimal ==> x \<noteq> 0" |
|
543 |
by auto |
|
544 |
||
| 15229 | 545 |
lemma Infinitesimal_hrabs_iff [iff]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
546 |
"(abs (x::hypreal) \<in> Infinitesimal) = (x \<in> Infinitesimal)" |
| 15003 | 547 |
by (auto simp add: abs_if) |
| 14370 | 548 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
549 |
lemma HFinite_diff_Infinitesimal_hrabs: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
550 |
"(x::hypreal) \<in> HFinite - Infinitesimal ==> abs x \<in> HFinite - Infinitesimal" |
| 14370 | 551 |
by blast |
552 |
||
553 |
lemma hrabs_less_Infinitesimal: |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
554 |
"[| e \<in> Infinitesimal; abs (x::hypreal) < e |] ==> x \<in> Infinitesimal" |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
555 |
by (auto simp add: Infinitesimal_def abs_less_iff) |
| 14370 | 556 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
557 |
lemma hrabs_le_Infinitesimal: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
558 |
"[| e \<in> Infinitesimal; abs (x::hypreal) \<le> e |] ==> x \<in> Infinitesimal" |
| 14370 | 559 |
by (blast dest: order_le_imp_less_or_eq intro: hrabs_less_Infinitesimal) |
560 |
||
561 |
lemma Infinitesimal_interval: |
|
562 |
"[| e \<in> Infinitesimal; e' \<in> Infinitesimal; e' < x ; x < e |] |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
563 |
==> (x::hypreal) \<in> Infinitesimal" |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
564 |
by (auto simp add: Infinitesimal_def abs_less_iff) |
| 14370 | 565 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
566 |
lemma Infinitesimal_interval2: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
567 |
"[| e \<in> Infinitesimal; e' \<in> Infinitesimal; |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
568 |
e' \<le> x ; x \<le> e |] ==> (x::hypreal) \<in> Infinitesimal" |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
569 |
by (auto intro: Infinitesimal_interval simp add: order_le_less) |
| 14370 | 570 |
|
571 |
lemma not_Infinitesimal_mult: |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
572 |
fixes x y :: "'a::real_normed_div_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
573 |
shows "[| x \<notin> Infinitesimal; y \<notin> Infinitesimal|] ==> (x*y) \<notin>Infinitesimal" |
| 20407 | 574 |
apply (unfold Infinitesimal_def, clarify, rename_tac r s) |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
575 |
apply (simp only: linorder_not_less hnorm_mult) |
| 20407 | 576 |
apply (drule_tac x = "r * s" in bspec) |
577 |
apply (fast intro: SReal_mult) |
|
578 |
apply (drule mp, blast intro: mult_pos_pos) |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
579 |
apply (drule_tac c = s and d = "hnorm y" and a = r and b = "hnorm x" in mult_mono) |
| 20407 | 580 |
apply (simp_all (no_asm_simp)) |
| 14370 | 581 |
done |
582 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
583 |
lemma Infinitesimal_mult_disj: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
584 |
fixes x y :: "'a::real_normed_div_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
585 |
shows "x*y \<in> Infinitesimal ==> x \<in> Infinitesimal | y \<in> Infinitesimal" |
| 14370 | 586 |
apply (rule ccontr) |
587 |
apply (drule de_Morgan_disj [THEN iffD1]) |
|
588 |
apply (fast dest: not_Infinitesimal_mult) |
|
589 |
done |
|
590 |
||
591 |
lemma HFinite_Infinitesimal_not_zero: "x \<in> HFinite-Infinitesimal ==> x \<noteq> 0" |
|
592 |
by blast |
|
593 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
594 |
lemma HFinite_Infinitesimal_diff_mult: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
595 |
fixes x y :: "'a::real_normed_div_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
596 |
shows "[| x \<in> HFinite - Infinitesimal; |
| 14370 | 597 |
y \<in> HFinite - Infinitesimal |
598 |
|] ==> (x*y) \<in> HFinite - Infinitesimal" |
|
599 |
apply clarify |
|
600 |
apply (blast dest: HFinite_mult not_Infinitesimal_mult) |
|
601 |
done |
|
602 |
||
603 |
lemma Infinitesimal_subset_HFinite: |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
604 |
"Infinitesimal \<subseteq> HFinite" |
| 14370 | 605 |
apply (simp add: Infinitesimal_def HFinite_def, auto) |
606 |
apply (rule_tac x = 1 in bexI, auto) |
|
607 |
done |
|
608 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
609 |
lemma Infinitesimal_hypreal_of_real_mult: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
610 |
"x \<in> Infinitesimal ==> x * hypreal_of_real r \<in> Infinitesimal" |
| 14370 | 611 |
by (erule HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult]) |
612 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
613 |
lemma Infinitesimal_hypreal_of_real_mult2: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
614 |
"x \<in> Infinitesimal ==> hypreal_of_real r * x \<in> Infinitesimal" |
| 14370 | 615 |
by (erule HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult2]) |
616 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
617 |
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
618 |
subsection{*The Infinitely Close Relation*}
|
| 14370 | 619 |
|
620 |
lemma mem_infmal_iff: "(x \<in> Infinitesimal) = (x @= 0)" |
|
621 |
by (simp add: Infinitesimal_def approx_def) |
|
622 |
||
| 20563 | 623 |
lemma approx_minus_iff: " (x @= y) = (x - y @= 0)" |
| 14370 | 624 |
by (simp add: approx_def) |
625 |
||
626 |
lemma approx_minus_iff2: " (x @= y) = (-y + x @= 0)" |
|
| 20563 | 627 |
by (simp add: approx_def diff_minus add_commute) |
| 14370 | 628 |
|
| 15229 | 629 |
lemma approx_refl [iff]: "x @= x" |
| 14370 | 630 |
by (simp add: approx_def Infinitesimal_def) |
631 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
632 |
lemma hypreal_minus_distrib1: "-(y + -(x::'a::ab_group_add)) = x + -y" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
633 |
by (simp add: add_commute) |
| 14477 | 634 |
|
| 14370 | 635 |
lemma approx_sym: "x @= y ==> y @= x" |
636 |
apply (simp add: approx_def) |
|
| 20563 | 637 |
apply (drule Infinitesimal_minus_iff [THEN iffD2]) |
638 |
apply simp |
|
| 14370 | 639 |
done |
640 |
||
641 |
lemma approx_trans: "[| x @= y; y @= z |] ==> x @= z" |
|
642 |
apply (simp add: approx_def) |
|
| 20563 | 643 |
apply (drule (1) Infinitesimal_add) |
644 |
apply (simp add: diff_def) |
|
| 14370 | 645 |
done |
646 |
||
647 |
lemma approx_trans2: "[| r @= x; s @= x |] ==> r @= s" |
|
648 |
by (blast intro: approx_sym approx_trans) |
|
649 |
||
650 |
lemma approx_trans3: "[| x @= r; x @= s|] ==> r @= s" |
|
651 |
by (blast intro: approx_sym approx_trans) |
|
652 |
||
653 |
lemma number_of_approx_reorient: "(number_of w @= x) = (x @= number_of w)" |
|
654 |
by (blast intro: approx_sym) |
|
655 |
||
656 |
lemma zero_approx_reorient: "(0 @= x) = (x @= 0)" |
|
657 |
by (blast intro: approx_sym) |
|
658 |
||
659 |
lemma one_approx_reorient: "(1 @= x) = (x @= 1)" |
|
660 |
by (blast intro: approx_sym) |
|
| 10751 | 661 |
|
662 |
||
| 19765 | 663 |
ML {*
|
664 |
local |
|
| 14370 | 665 |
(*** re-orientation, following HOL/Integ/Bin.ML |
666 |
We re-orient x @=y where x is 0, 1 or a numeral, unless y is as well! |
|
667 |
***) |
|
668 |
||
669 |
(*reorientation simprules using ==, for the following simproc*) |
|
| 19765 | 670 |
val meta_zero_approx_reorient = thm "zero_approx_reorient" RS eq_reflection; |
671 |
val meta_one_approx_reorient = thm "one_approx_reorient" RS eq_reflection; |
|
672 |
val meta_number_of_approx_reorient = thm "number_of_approx_reorient" RS eq_reflection |
|
| 14370 | 673 |
|
674 |
(*reorientation simplification procedure: reorients (polymorphic) |
|
675 |
0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*) |
|
676 |
fun reorient_proc sg _ (_ $ t $ u) = |
|
677 |
case u of |
|
| 15531 | 678 |
Const("0", _) => NONE
|
679 |
| Const("1", _) => NONE
|
|
680 |
| Const("Numeral.number_of", _) $ _ => NONE
|
|
681 |
| _ => SOME (case t of |
|
| 14370 | 682 |
Const("0", _) => meta_zero_approx_reorient
|
683 |
| Const("1", _) => meta_one_approx_reorient
|
|
684 |
| Const("Numeral.number_of", _) $ _ =>
|
|
685 |
meta_number_of_approx_reorient); |
|
686 |
||
| 19765 | 687 |
in |
| 14370 | 688 |
val approx_reorient_simproc = |
| 20485 | 689 |
Int_Numeral_Base_Simprocs.prep_simproc |
| 14370 | 690 |
("reorient_simproc", ["0@=x", "1@=x", "number_of w @= x"], reorient_proc);
|
| 19765 | 691 |
end; |
| 14370 | 692 |
|
693 |
Addsimprocs [approx_reorient_simproc]; |
|
694 |
*} |
|
695 |
||
696 |
lemma Infinitesimal_approx_minus: "(x-y \<in> Infinitesimal) = (x @= y)" |
|
| 20563 | 697 |
by (simp add: approx_minus_iff [symmetric] mem_infmal_iff) |
| 14370 | 698 |
|
699 |
lemma approx_monad_iff: "(x @= y) = (monad(x)=monad(y))" |
|
700 |
apply (simp add: monad_def) |
|
701 |
apply (auto dest: approx_sym elim!: approx_trans equalityCE) |
|
702 |
done |
|
703 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
704 |
lemma Infinitesimal_approx: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
705 |
"[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> x @= y" |
| 14370 | 706 |
apply (simp add: mem_infmal_iff) |
707 |
apply (blast intro: approx_trans approx_sym) |
|
708 |
done |
|
709 |
||
710 |
lemma approx_add: "[| a @= b; c @= d |] ==> a+c @= b+d" |
|
711 |
proof (unfold approx_def) |
|
| 20563 | 712 |
assume inf: "a - b \<in> Infinitesimal" "c - d \<in> Infinitesimal" |
713 |
have "a + c - (b + d) = (a - b) + (c - d)" by simp |
|
| 14370 | 714 |
also have "... \<in> Infinitesimal" using inf by (rule Infinitesimal_add) |
| 20563 | 715 |
finally show "a + c - (b + d) \<in> Infinitesimal" . |
| 14370 | 716 |
qed |
717 |
||
718 |
lemma approx_minus: "a @= b ==> -a @= -b" |
|
719 |
apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym]) |
|
720 |
apply (drule approx_minus_iff [THEN iffD1]) |
|
| 20563 | 721 |
apply (simp add: add_commute diff_def) |
| 14370 | 722 |
done |
723 |
||
724 |
lemma approx_minus2: "-a @= -b ==> a @= b" |
|
725 |
by (auto dest: approx_minus) |
|
726 |
||
| 15229 | 727 |
lemma approx_minus_cancel [simp]: "(-a @= -b) = (a @= b)" |
| 14370 | 728 |
by (blast intro: approx_minus approx_minus2) |
729 |
||
730 |
lemma approx_add_minus: "[| a @= b; c @= d |] ==> a + -c @= b + -d" |
|
731 |
by (blast intro!: approx_add approx_minus) |
|
732 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
733 |
lemma approx_mult1: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
734 |
fixes a b c :: "'a::real_normed_algebra star" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
735 |
shows "[| a @= b; c: HFinite|] ==> a*c @= b*c" |
| 20563 | 736 |
by (simp add: approx_def Infinitesimal_HFinite_mult |
737 |
left_diff_distrib [symmetric]) |
|
| 14370 | 738 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
739 |
lemma approx_mult2: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
740 |
fixes a b c :: "'a::real_normed_algebra star" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
741 |
shows "[|a @= b; c: HFinite|] ==> c*a @= c*b" |
| 20563 | 742 |
by (simp add: approx_def Infinitesimal_HFinite_mult2 |
743 |
right_diff_distrib [symmetric]) |
|
| 14370 | 744 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
745 |
lemma approx_mult_subst: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
746 |
fixes u v x y :: "'a::real_normed_algebra star" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
747 |
shows "[|u @= v*x; x @= y; v \<in> HFinite|] ==> u @= v*y" |
| 14370 | 748 |
by (blast intro: approx_mult2 approx_trans) |
749 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
750 |
lemma approx_mult_subst2: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
751 |
fixes u v x y :: "'a::real_normed_algebra star" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
752 |
shows "[| u @= x*v; x @= y; v \<in> HFinite |] ==> u @= y*v" |
| 14370 | 753 |
by (blast intro: approx_mult1 approx_trans) |
754 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
755 |
lemma approx_mult_subst_star_of: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
756 |
fixes u x y :: "'a::real_normed_algebra star" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
757 |
shows "[| u @= x*star_of v; x @= y |] ==> u @= y*star_of v" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
758 |
by (auto intro: approx_mult_subst2) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
759 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
760 |
lemma approx_mult_subst_SReal: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
761 |
"[| u @= x*hypreal_of_real v; x @= y |] ==> u @= y*hypreal_of_real v" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
762 |
by (rule approx_mult_subst_star_of) |
| 14370 | 763 |
|
764 |
lemma approx_eq_imp: "a = b ==> a @= b" |
|
765 |
by (simp add: approx_def) |
|
766 |
||
767 |
lemma Infinitesimal_minus_approx: "x \<in> Infinitesimal ==> -x @= x" |
|
768 |
by (blast intro: Infinitesimal_minus_iff [THEN iffD2] |
|
769 |
mem_infmal_iff [THEN iffD1] approx_trans2) |
|
770 |
||
| 20563 | 771 |
lemma bex_Infinitesimal_iff: "(\<exists>y \<in> Infinitesimal. x - z = y) = (x @= z)" |
| 14370 | 772 |
by (simp add: approx_def) |
773 |
||
774 |
lemma bex_Infinitesimal_iff2: "(\<exists>y \<in> Infinitesimal. x = z + y) = (x @= z)" |
|
775 |
by (force simp add: bex_Infinitesimal_iff [symmetric]) |
|
776 |
||
777 |
lemma Infinitesimal_add_approx: "[| y \<in> Infinitesimal; x + y = z |] ==> x @= z" |
|
778 |
apply (rule bex_Infinitesimal_iff [THEN iffD1]) |
|
779 |
apply (drule Infinitesimal_minus_iff [THEN iffD2]) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
780 |
apply (auto simp add: add_assoc [symmetric]) |
| 14370 | 781 |
done |
782 |
||
783 |
lemma Infinitesimal_add_approx_self: "y \<in> Infinitesimal ==> x @= x + y" |
|
784 |
apply (rule bex_Infinitesimal_iff [THEN iffD1]) |
|
785 |
apply (drule Infinitesimal_minus_iff [THEN iffD2]) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
786 |
apply (auto simp add: add_assoc [symmetric]) |
| 14370 | 787 |
done |
788 |
||
789 |
lemma Infinitesimal_add_approx_self2: "y \<in> Infinitesimal ==> x @= y + x" |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
790 |
by (auto dest: Infinitesimal_add_approx_self simp add: add_commute) |
| 14370 | 791 |
|
792 |
lemma Infinitesimal_add_minus_approx_self: "y \<in> Infinitesimal ==> x @= x + -y" |
|
793 |
by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2]) |
|
794 |
||
795 |
lemma Infinitesimal_add_cancel: "[| y \<in> Infinitesimal; x+y @= z|] ==> x @= z" |
|
796 |
apply (drule_tac x = x in Infinitesimal_add_approx_self [THEN approx_sym]) |
|
797 |
apply (erule approx_trans3 [THEN approx_sym], assumption) |
|
798 |
done |
|
799 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
800 |
lemma Infinitesimal_add_right_cancel: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
801 |
"[| y \<in> Infinitesimal; x @= z + y|] ==> x @= z" |
| 14370 | 802 |
apply (drule_tac x = z in Infinitesimal_add_approx_self2 [THEN approx_sym]) |
803 |
apply (erule approx_trans3 [THEN approx_sym]) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
804 |
apply (simp add: add_commute) |
| 14370 | 805 |
apply (erule approx_sym) |
806 |
done |
|
807 |
||
808 |
lemma approx_add_left_cancel: "d + b @= d + c ==> b @= c" |
|
809 |
apply (drule approx_minus_iff [THEN iffD1]) |
|
| 15539 | 810 |
apply (simp add: approx_minus_iff [symmetric] add_ac) |
| 14370 | 811 |
done |
812 |
||
813 |
lemma approx_add_right_cancel: "b + d @= c + d ==> b @= c" |
|
814 |
apply (rule approx_add_left_cancel) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
815 |
apply (simp add: add_commute) |
| 14370 | 816 |
done |
817 |
||
818 |
lemma approx_add_mono1: "b @= c ==> d + b @= d + c" |
|
819 |
apply (rule approx_minus_iff [THEN iffD2]) |
|
| 15539 | 820 |
apply (simp add: approx_minus_iff [symmetric] add_ac) |
| 14370 | 821 |
done |
822 |
||
823 |
lemma approx_add_mono2: "b @= c ==> b + a @= c + a" |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
824 |
by (simp add: add_commute approx_add_mono1) |
| 14370 | 825 |
|
| 15229 | 826 |
lemma approx_add_left_iff [simp]: "(a + b @= a + c) = (b @= c)" |
| 14370 | 827 |
by (fast elim: approx_add_left_cancel approx_add_mono1) |
828 |
||
| 15229 | 829 |
lemma approx_add_right_iff [simp]: "(b + a @= c + a) = (b @= c)" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
830 |
by (simp add: add_commute) |
| 14370 | 831 |
|
832 |
lemma approx_HFinite: "[| x \<in> HFinite; x @= y |] ==> y \<in> HFinite" |
|
833 |
apply (drule bex_Infinitesimal_iff2 [THEN iffD2], safe) |
|
834 |
apply (drule Infinitesimal_subset_HFinite [THEN subsetD, THEN HFinite_minus_iff [THEN iffD2]]) |
|
835 |
apply (drule HFinite_add) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
836 |
apply (auto simp add: add_assoc) |
| 14370 | 837 |
done |
838 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
839 |
lemma approx_star_of_HFinite: "x @= star_of D ==> x \<in> HFinite" |
| 14370 | 840 |
by (rule approx_sym [THEN [2] approx_HFinite], auto) |
841 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
842 |
lemma approx_hypreal_of_real_HFinite: "x @= hypreal_of_real D ==> x \<in> HFinite" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
843 |
by (rule approx_star_of_HFinite) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
844 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
845 |
lemma approx_mult_HFinite: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
846 |
fixes a b c d :: "'a::real_normed_algebra star" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
847 |
shows "[|a @= b; c @= d; b: HFinite; d: HFinite|] ==> a*c @= b*d" |
| 14370 | 848 |
apply (rule approx_trans) |
849 |
apply (rule_tac [2] approx_mult2) |
|
850 |
apply (rule approx_mult1) |
|
851 |
prefer 2 apply (blast intro: approx_HFinite approx_sym, auto) |
|
852 |
done |
|
853 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
854 |
lemma approx_mult_star_of: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
855 |
fixes a c :: "'a::real_normed_algebra star" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
856 |
shows "[|a @= star_of b; c @= star_of d |] |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
857 |
==> a*c @= star_of b*star_of d" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
858 |
by (blast intro!: approx_mult_HFinite approx_star_of_HFinite HFinite_star_of) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
859 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
860 |
lemma approx_mult_hypreal_of_real: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
861 |
"[|a @= hypreal_of_real b; c @= hypreal_of_real d |] |
| 14370 | 862 |
==> a*c @= hypreal_of_real b*hypreal_of_real d" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
863 |
by (rule approx_mult_star_of) |
| 14370 | 864 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
865 |
lemma approx_SReal_mult_cancel_zero: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
866 |
"[| (a::hypreal) \<in> Reals; a \<noteq> 0; a*x @= 0 |] ==> x @= 0" |
| 14370 | 867 |
apply (drule SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
868 |
apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) |
| 14370 | 869 |
done |
870 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
871 |
lemma approx_mult_SReal1: "[| (a::hypreal) \<in> Reals; x @= 0 |] ==> x*a @= 0" |
| 14370 | 872 |
by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1) |
873 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
874 |
lemma approx_mult_SReal2: "[| (a::hypreal) \<in> Reals; x @= 0 |] ==> a*x @= 0" |
| 14370 | 875 |
by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2) |
876 |
||
| 15229 | 877 |
lemma approx_mult_SReal_zero_cancel_iff [simp]: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
878 |
"[|(a::hypreal) \<in> Reals; a \<noteq> 0 |] ==> (a*x @= 0) = (x @= 0)" |
| 14370 | 879 |
by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2) |
880 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
881 |
lemma approx_SReal_mult_cancel: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
882 |
"[| (a::hypreal) \<in> Reals; a \<noteq> 0; a* w @= a*z |] ==> w @= z" |
| 14370 | 883 |
apply (drule SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
884 |
apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) |
| 14370 | 885 |
done |
886 |
||
| 15229 | 887 |
lemma approx_SReal_mult_cancel_iff1 [simp]: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
888 |
"[| (a::hypreal) \<in> Reals; a \<noteq> 0|] ==> (a* w @= a*z) = (w @= z)" |
| 15229 | 889 |
by (auto intro!: approx_mult2 SReal_subset_HFinite [THEN subsetD] |
890 |
intro: approx_SReal_mult_cancel) |
|
| 14370 | 891 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
892 |
lemma approx_le_bound: "[| (z::hypreal) \<le> f; f @= g; g \<le> z |] ==> f @= z" |
| 14370 | 893 |
apply (simp add: bex_Infinitesimal_iff2 [symmetric], auto) |
894 |
apply (rule_tac x = "g+y-z" in bexI) |
|
895 |
apply (simp (no_asm)) |
|
896 |
apply (rule Infinitesimal_interval2) |
|
897 |
apply (rule_tac [2] Infinitesimal_zero, auto) |
|
898 |
done |
|
899 |
||
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
900 |
lemma approx_hnorm: |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
901 |
fixes x y :: "'a::real_normed_vector star" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
902 |
shows "x \<approx> y \<Longrightarrow> hnorm x \<approx> hnorm y" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
903 |
proof (unfold approx_def) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
904 |
assume "x - y \<in> Infinitesimal" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
905 |
hence 1: "hnorm (x - y) \<in> Infinitesimal" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
906 |
by (simp only: Infinitesimal_hnorm_iff) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
907 |
moreover have 2: "(0::real star) \<in> Infinitesimal" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
908 |
by (rule Infinitesimal_zero) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
909 |
moreover have 3: "0 \<le> \<bar>hnorm x - hnorm y\<bar>" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
910 |
by (rule abs_ge_zero) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
911 |
moreover have 4: "\<bar>hnorm x - hnorm y\<bar> \<le> hnorm (x - y)" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
912 |
by (rule hnorm_triangle_ineq3) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
913 |
ultimately have "\<bar>hnorm x - hnorm y\<bar> \<in> Infinitesimal" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
914 |
by (rule Infinitesimal_interval2) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
915 |
thus "hnorm x - hnorm y \<in> Infinitesimal" |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
916 |
by (simp only: Infinitesimal_hrabs_iff) |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
917 |
qed |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
918 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
919 |
|
| 15539 | 920 |
subsection{* Zero is the Only Infinitesimal that is also a Real*}
|
| 14370 | 921 |
|
922 |
lemma Infinitesimal_less_SReal: |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
923 |
"[| (x::hypreal) \<in> Reals; y \<in> Infinitesimal; 0 < x |] ==> y < x" |
| 14370 | 924 |
apply (simp add: Infinitesimal_def) |
925 |
apply (rule abs_ge_self [THEN order_le_less_trans], auto) |
|
926 |
done |
|
927 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
928 |
lemma Infinitesimal_less_SReal2: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
929 |
"(y::hypreal) \<in> Infinitesimal ==> \<forall>r \<in> Reals. 0 < r --> y < r" |
| 14370 | 930 |
by (blast intro: Infinitesimal_less_SReal) |
931 |
||
932 |
lemma SReal_not_Infinitesimal: |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
933 |
"[| 0 < y; (y::hypreal) \<in> Reals|] ==> y \<notin> Infinitesimal" |
| 14370 | 934 |
apply (simp add: Infinitesimal_def) |
| 15003 | 935 |
apply (auto simp add: abs_if) |
| 14370 | 936 |
done |
937 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
938 |
lemma SReal_minus_not_Infinitesimal: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
939 |
"[| y < 0; (y::hypreal) \<in> Reals |] ==> y \<notin> Infinitesimal" |
| 14370 | 940 |
apply (subst Infinitesimal_minus_iff [symmetric]) |
941 |
apply (rule SReal_not_Infinitesimal, auto) |
|
942 |
done |
|
943 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
944 |
lemma SReal_Int_Infinitesimal_zero: "Reals Int Infinitesimal = {0::hypreal}"
|
| 14370 | 945 |
apply auto |
946 |
apply (cut_tac x = x and y = 0 in linorder_less_linear) |
|
947 |
apply (blast dest: SReal_not_Infinitesimal SReal_minus_not_Infinitesimal) |
|
948 |
done |
|
949 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
950 |
lemma SReal_Infinitesimal_zero: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
951 |
"[| (x::hypreal) \<in> Reals; x \<in> Infinitesimal|] ==> x = 0" |
| 14370 | 952 |
by (cut_tac SReal_Int_Infinitesimal_zero, blast) |
953 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
954 |
lemma SReal_HFinite_diff_Infinitesimal: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
955 |
"[| (x::hypreal) \<in> Reals; x \<noteq> 0 |] ==> x \<in> HFinite - Infinitesimal" |
| 14370 | 956 |
by (auto dest: SReal_Infinitesimal_zero SReal_subset_HFinite [THEN subsetD]) |
957 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
958 |
lemma hypreal_of_real_HFinite_diff_Infinitesimal: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
959 |
"hypreal_of_real x \<noteq> 0 ==> hypreal_of_real x \<in> HFinite - Infinitesimal" |
| 14370 | 960 |
by (rule SReal_HFinite_diff_Infinitesimal, auto) |
961 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
962 |
lemma star_of_Infinitesimal_iff_0 [iff]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
963 |
"(star_of x \<in> Infinitesimal) = (x = 0)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
964 |
apply (auto simp add: Infinitesimal_def) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
965 |
apply (drule_tac x="hnorm (star_of x)" in bspec) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
966 |
apply (simp add: hnorm_def) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
967 |
apply simp |
| 14370 | 968 |
done |
969 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
970 |
lemma star_of_HFinite_diff_Infinitesimal: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
971 |
"x \<noteq> 0 ==> star_of x \<in> HFinite - Infinitesimal" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
972 |
by simp |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
973 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
974 |
lemma hypreal_of_real_Infinitesimal_iff_0: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
975 |
"(hypreal_of_real x \<in> Infinitesimal) = (x=0)" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
976 |
by (rule star_of_Infinitesimal_iff_0) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
977 |
|
| 15229 | 978 |
lemma number_of_not_Infinitesimal [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
979 |
"number_of w \<noteq> (0::hypreal) ==> (number_of w :: hypreal) \<notin> Infinitesimal" |
| 14370 | 980 |
by (fast dest: SReal_number_of [THEN SReal_Infinitesimal_zero]) |
981 |
||
982 |
(*again: 1 is a special case, but not 0 this time*) |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
983 |
lemma one_not_Infinitesimal [simp]: |
| 20633 | 984 |
"(1::'a::{real_normed_vector,zero_neq_one} star) \<notin> Infinitesimal"
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
985 |
apply (simp only: star_one_def star_of_Infinitesimal_iff_0) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
986 |
apply simp |
| 14370 | 987 |
done |
988 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
989 |
lemma approx_SReal_not_zero: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
990 |
"[| (y::hypreal) \<in> Reals; x @= y; y\<noteq> 0 |] ==> x \<noteq> 0" |
| 14370 | 991 |
apply (cut_tac x = 0 and y = y in linorder_less_linear, simp) |
992 |
apply (blast dest: approx_sym [THEN mem_infmal_iff [THEN iffD2]] SReal_not_Infinitesimal SReal_minus_not_Infinitesimal) |
|
993 |
done |
|
994 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
995 |
lemma HFinite_diff_Infinitesimal_approx: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
996 |
"[| x @= y; y \<in> HFinite - Infinitesimal |] |
| 14370 | 997 |
==> x \<in> HFinite - Infinitesimal" |
998 |
apply (auto intro: approx_sym [THEN [2] approx_HFinite] |
|
999 |
simp add: mem_infmal_iff) |
|
1000 |
apply (drule approx_trans3, assumption) |
|
1001 |
apply (blast dest: approx_sym) |
|
1002 |
done |
|
1003 |
||
1004 |
(*The premise y\<noteq>0 is essential; otherwise x/y =0 and we lose the |
|
1005 |
HFinite premise.*) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1006 |
lemma Infinitesimal_ratio: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1007 |
fixes x y :: "'a::{real_normed_div_algebra,field} star"
|
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1008 |
shows "[| y \<noteq> 0; y \<in> Infinitesimal; x/y \<in> HFinite |] |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1009 |
==> x \<in> Infinitesimal" |
| 14370 | 1010 |
apply (drule Infinitesimal_HFinite_mult2, assumption) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1011 |
apply (simp add: divide_inverse mult_assoc) |
| 14370 | 1012 |
done |
1013 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1014 |
lemma Infinitesimal_SReal_divide: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1015 |
"[| (x::hypreal) \<in> Infinitesimal; y \<in> Reals |] ==> x/y \<in> Infinitesimal" |
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14420
diff
changeset
|
1016 |
apply (simp add: divide_inverse) |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1017 |
apply (auto intro!: Infinitesimal_HFinite_mult |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1018 |
dest!: SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1019 |
done |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1020 |
|
| 14370 | 1021 |
(*------------------------------------------------------------------ |
1022 |
Standard Part Theorem: Every finite x: R* is infinitely |
|
1023 |
close to a unique real number (i.e a member of Reals) |
|
1024 |
------------------------------------------------------------------*) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1025 |
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1026 |
subsection{* Uniqueness: Two Infinitely Close Reals are Equal*}
|
| 14370 | 1027 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1028 |
lemma star_of_approx_iff [simp]: "(star_of x @= star_of y) = (x = y)" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1029 |
apply safe |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1030 |
apply (simp add: approx_def) |
| 20563 | 1031 |
apply (simp only: star_of_diff [symmetric]) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1032 |
apply (simp only: star_of_Infinitesimal_iff_0) |
| 20563 | 1033 |
apply simp |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1034 |
done |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1035 |
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1036 |
lemma SReal_approx_iff: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1037 |
"[|(x::hypreal) \<in> Reals; y \<in> Reals|] ==> (x @= y) = (x = y)" |
| 14370 | 1038 |
apply auto |
1039 |
apply (simp add: approx_def) |
|
| 20563 | 1040 |
apply (drule (1) SReal_diff) |
1041 |
apply (drule (1) SReal_Infinitesimal_zero) |
|
1042 |
apply simp |
|
| 14370 | 1043 |
done |
1044 |
||
| 15229 | 1045 |
lemma number_of_approx_iff [simp]: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1046 |
"(number_of v @= (number_of w :: 'a::{number,real_normed_vector} star)) =
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1047 |
(number_of v = (number_of w :: 'a))" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1048 |
apply (unfold star_number_def) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1049 |
apply (rule star_of_approx_iff) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1050 |
done |
| 14370 | 1051 |
|
1052 |
(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*) |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1053 |
lemma [simp]: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1054 |
"(number_of w @= (0::'a::{number,real_normed_vector} star)) =
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1055 |
(number_of w = (0::'a))" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1056 |
"((0::'a::{number,real_normed_vector} star) @= number_of w) =
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1057 |
(number_of w = (0::'a))" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1058 |
"(number_of w @= (1::'b::{number,one,real_normed_vector} star)) =
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1059 |
(number_of w = (1::'b))" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1060 |
"((1::'b::{number,one,real_normed_vector} star) @= number_of w) =
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1061 |
(number_of w = (1::'b))" |
| 20633 | 1062 |
"~ (0 @= (1::'c::{zero_neq_one,real_normed_vector} star))"
|
1063 |
"~ (1 @= (0::'c::{zero_neq_one,real_normed_vector} star))"
|
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1064 |
apply (unfold star_number_def star_zero_def star_one_def) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1065 |
apply (unfold star_of_approx_iff) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1066 |
by (auto intro: sym) |
| 14370 | 1067 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1068 |
lemma hypreal_of_real_approx_iff: |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1069 |
"(hypreal_of_real k @= hypreal_of_real m) = (k = m)" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1070 |
by (rule star_of_approx_iff) |
| 14370 | 1071 |
|
| 15229 | 1072 |
lemma hypreal_of_real_approx_number_of_iff [simp]: |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1073 |
"(hypreal_of_real k @= number_of w) = (k = number_of w)" |
| 14370 | 1074 |
by (subst hypreal_of_real_approx_iff [symmetric], auto) |
1075 |
||
1076 |
(*And also for 0 and 1.*) |
|
1077 |
(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*) |
|
1078 |
lemma [simp]: "(hypreal_of_real k @= 0) = (k = 0)" |
|
1079 |
"(hypreal_of_real k @= 1) = (k = 1)" |
|
1080 |
by (simp_all add: hypreal_of_real_approx_iff [symmetric]) |
|
1081 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1082 |
lemma approx_unique_real: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1083 |
"[| (r::hypreal) \<in> Reals; s \<in> Reals; r @= x; s @= x|] ==> r = s" |
| 14370 | 1084 |
by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2) |
1085 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1086 |
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1087 |
subsection{* Existence of Unique Real Infinitely Close*}
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1088 |
|
| 14370 | 1089 |
(* lemma about lubs *) |
1090 |
lemma hypreal_isLub_unique: |
|
1091 |
"[| isLub R S x; isLub R S y |] ==> x = (y::hypreal)" |
|
1092 |
apply (frule isLub_isUb) |
|
1093 |
apply (frule_tac x = y in isLub_isUb) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1094 |
apply (blast intro!: order_antisym dest!: isLub_le_isUb) |
| 14370 | 1095 |
done |
1096 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1097 |
lemma lemma_st_part_ub: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1098 |
"(x::hypreal) \<in> HFinite ==> \<exists>u. isUb Reals {s. s \<in> Reals & s < x} u"
|
| 14370 | 1099 |
apply (drule HFiniteD, safe) |
1100 |
apply (rule exI, rule isUbI) |
|
1101 |
apply (auto intro: setleI isUbI simp add: abs_less_iff) |
|
1102 |
done |
|
1103 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1104 |
lemma lemma_st_part_nonempty: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1105 |
"(x::hypreal) \<in> HFinite ==> \<exists>y. y \<in> {s. s \<in> Reals & s < x}"
|
| 14370 | 1106 |
apply (drule HFiniteD, safe) |
1107 |
apply (drule SReal_minus) |
|
1108 |
apply (rule_tac x = "-t" in exI) |
|
1109 |
apply (auto simp add: abs_less_iff) |
|
1110 |
done |
|
1111 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1112 |
lemma lemma_st_part_subset: "{s. s \<in> Reals & s < x} \<subseteq> Reals"
|
| 14370 | 1113 |
by auto |
1114 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1115 |
lemma lemma_st_part_lub: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1116 |
"(x::hypreal) \<in> HFinite ==> \<exists>t. isLub Reals {s. s \<in> Reals & s < x} t"
|
| 14370 | 1117 |
by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty lemma_st_part_subset) |
1118 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1119 |
lemma lemma_hypreal_le_left_cancel: "((t::hypreal) + r \<le> t) = (r \<le> 0)" |
| 14370 | 1120 |
apply safe |
1121 |
apply (drule_tac c = "-t" in add_left_mono) |
|
1122 |
apply (drule_tac [2] c = t in add_left_mono) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1123 |
apply (auto simp add: add_assoc [symmetric]) |
| 14370 | 1124 |
done |
1125 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1126 |
lemma lemma_st_part_le1: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1127 |
"[| (x::hypreal) \<in> HFinite; isLub Reals {s. s \<in> Reals & s < x} t;
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1128 |
r \<in> Reals; 0 < r |] ==> x \<le> t + r" |
| 14370 | 1129 |
apply (frule isLubD1a) |
1130 |
apply (rule ccontr, drule linorder_not_le [THEN iffD2]) |
|
1131 |
apply (drule_tac x = t in SReal_add, assumption) |
|
1132 |
apply (drule_tac y = "t + r" in isLubD1 [THEN setleD], auto) |
|
1133 |
done |
|
1134 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1135 |
lemma hypreal_setle_less_trans: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1136 |
"[| S *<= (x::hypreal); x < y |] ==> S *<= y" |
| 14370 | 1137 |
apply (simp add: setle_def) |
1138 |
apply (auto dest!: bspec order_le_less_trans intro: order_less_imp_le) |
|
1139 |
done |
|
1140 |
||
1141 |
lemma hypreal_gt_isUb: |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1142 |
"[| isUb R S (x::hypreal); x < y; y \<in> R |] ==> isUb R S y" |
| 14370 | 1143 |
apply (simp add: isUb_def) |
1144 |
apply (blast intro: hypreal_setle_less_trans) |
|
1145 |
done |
|
1146 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1147 |
lemma lemma_st_part_gt_ub: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1148 |
"[| (x::hypreal) \<in> HFinite; x < y; y \<in> Reals |] |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1149 |
==> isUb Reals {s. s \<in> Reals & s < x} y"
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1150 |
by (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI) |
| 14370 | 1151 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1152 |
lemma lemma_minus_le_zero: "t \<le> t + -r ==> r \<le> (0::hypreal)" |
| 14370 | 1153 |
apply (drule_tac c = "-t" in add_left_mono) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1154 |
apply (auto simp add: add_assoc [symmetric]) |
| 14370 | 1155 |
done |
1156 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1157 |
lemma lemma_st_part_le2: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1158 |
"[| (x::hypreal) \<in> HFinite; |
| 14370 | 1159 |
isLub Reals {s. s \<in> Reals & s < x} t;
|
1160 |
r \<in> Reals; 0 < r |] |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1161 |
==> t + -r \<le> x" |
| 14370 | 1162 |
apply (frule isLubD1a) |
1163 |
apply (rule ccontr, drule linorder_not_le [THEN iffD1]) |
|
1164 |
apply (drule SReal_minus, drule_tac x = t in SReal_add, assumption) |
|
1165 |
apply (drule lemma_st_part_gt_ub, assumption+) |
|
1166 |
apply (drule isLub_le_isUb, assumption) |
|
1167 |
apply (drule lemma_minus_le_zero) |
|
1168 |
apply (auto dest: order_less_le_trans) |
|
1169 |
done |
|
1170 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1171 |
lemma lemma_st_part1a: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1172 |
"[| (x::hypreal) \<in> HFinite; |
| 14370 | 1173 |
isLub Reals {s. s \<in> Reals & s < x} t;
|
1174 |
r \<in> Reals; 0 < r |] |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1175 |
==> x + -t \<le> r" |
| 15229 | 1176 |
apply (subgoal_tac "x \<le> t+r") |
1177 |
apply (auto intro: lemma_st_part_le1) |
|
1178 |
done |
|
| 14370 | 1179 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1180 |
lemma lemma_st_part2a: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1181 |
"[| (x::hypreal) \<in> HFinite; |
| 14370 | 1182 |
isLub Reals {s. s \<in> Reals & s < x} t;
|
1183 |
r \<in> Reals; 0 < r |] |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1184 |
==> -(x + -t) \<le> r" |
| 15229 | 1185 |
apply (subgoal_tac "(t + -r \<le> x)") |
1186 |
apply (auto intro: lemma_st_part_le2) |
|
| 14370 | 1187 |
done |
1188 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1189 |
lemma lemma_SReal_ub: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1190 |
"(x::hypreal) \<in> Reals ==> isUb Reals {s. s \<in> Reals & s < x} x"
|
| 14370 | 1191 |
by (auto intro: isUbI setleI order_less_imp_le) |
1192 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1193 |
lemma lemma_SReal_lub: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1194 |
"(x::hypreal) \<in> Reals ==> isLub Reals {s. s \<in> Reals & s < x} x"
|
| 14370 | 1195 |
apply (auto intro!: isLubI2 lemma_SReal_ub setgeI) |
1196 |
apply (frule isUbD2a) |
|
1197 |
apply (rule_tac x = x and y = y in linorder_cases) |
|
1198 |
apply (auto intro!: order_less_imp_le) |
|
1199 |
apply (drule SReal_dense, assumption, assumption, safe) |
|
1200 |
apply (drule_tac y = r in isUbD) |
|
1201 |
apply (auto dest: order_less_le_trans) |
|
1202 |
done |
|
1203 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1204 |
lemma lemma_st_part_not_eq1: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1205 |
"[| (x::hypreal) \<in> HFinite; |
| 14370 | 1206 |
isLub Reals {s. s \<in> Reals & s < x} t;
|
1207 |
r \<in> Reals; 0 < r |] |
|
1208 |
==> x + -t \<noteq> r" |
|
1209 |
apply auto |
|
1210 |
apply (frule isLubD1a [THEN SReal_minus]) |
|
1211 |
apply (drule SReal_add_cancel, assumption) |
|
1212 |
apply (drule_tac x = x in lemma_SReal_lub) |
|
1213 |
apply (drule hypreal_isLub_unique, assumption, auto) |
|
1214 |
done |
|
1215 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1216 |
lemma lemma_st_part_not_eq2: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1217 |
"[| (x::hypreal) \<in> HFinite; |
| 14370 | 1218 |
isLub Reals {s. s \<in> Reals & s < x} t;
|
1219 |
r \<in> Reals; 0 < r |] |
|
1220 |
==> -(x + -t) \<noteq> r" |
|
| 15539 | 1221 |
apply (auto) |
| 14370 | 1222 |
apply (frule isLubD1a) |
1223 |
apply (drule SReal_add_cancel, assumption) |
|
1224 |
apply (drule_tac x = "-x" in SReal_minus, simp) |
|
1225 |
apply (drule_tac x = x in lemma_SReal_lub) |
|
1226 |
apply (drule hypreal_isLub_unique, assumption, auto) |
|
1227 |
done |
|
1228 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1229 |
lemma lemma_st_part_major: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1230 |
"[| (x::hypreal) \<in> HFinite; |
| 14370 | 1231 |
isLub Reals {s. s \<in> Reals & s < x} t;
|
1232 |
r \<in> Reals; 0 < r |] |
|
| 20563 | 1233 |
==> abs (x - t) < r" |
| 14370 | 1234 |
apply (frule lemma_st_part1a) |
1235 |
apply (frule_tac [4] lemma_st_part2a, auto) |
|
1236 |
apply (drule order_le_imp_less_or_eq)+ |
|
1237 |
apply (auto dest: lemma_st_part_not_eq1 lemma_st_part_not_eq2 simp add: abs_less_iff) |
|
1238 |
done |
|
1239 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1240 |
lemma lemma_st_part_major2: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1241 |
"[| (x::hypreal) \<in> HFinite; isLub Reals {s. s \<in> Reals & s < x} t |]
|
| 20563 | 1242 |
==> \<forall>r \<in> Reals. 0 < r --> abs (x - t) < r" |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1243 |
by (blast dest!: lemma_st_part_major) |
| 14370 | 1244 |
|
| 15229 | 1245 |
|
1246 |
text{*Existence of real and Standard Part Theorem*}
|
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1247 |
lemma lemma_st_part_Ex: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1248 |
"(x::hypreal) \<in> HFinite |
| 20563 | 1249 |
==> \<exists>t \<in> Reals. \<forall>r \<in> Reals. 0 < r --> abs (x - t) < r" |
| 14370 | 1250 |
apply (frule lemma_st_part_lub, safe) |
1251 |
apply (frule isLubD1a) |
|
1252 |
apply (blast dest: lemma_st_part_major2) |
|
1253 |
done |
|
1254 |
||
1255 |
lemma st_part_Ex: |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1256 |
"(x::hypreal) \<in> HFinite ==> \<exists>t \<in> Reals. x @= t" |
| 14370 | 1257 |
apply (simp add: approx_def Infinitesimal_def) |
1258 |
apply (drule lemma_st_part_Ex, auto) |
|
1259 |
done |
|
1260 |
||
| 15229 | 1261 |
text{*There is a unique real infinitely close*}
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1262 |
lemma st_part_Ex1: "x \<in> HFinite ==> EX! t::hypreal. t \<in> Reals & x @= t" |
| 14370 | 1263 |
apply (drule st_part_Ex, safe) |
1264 |
apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym) |
|
1265 |
apply (auto intro!: approx_unique_real) |
|
1266 |
done |
|
1267 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1268 |
subsection{* Finite, Infinite and Infinitesimal*}
|
| 14370 | 1269 |
|
| 15229 | 1270 |
lemma HFinite_Int_HInfinite_empty [simp]: "HFinite Int HInfinite = {}"
|
| 14370 | 1271 |
apply (simp add: HFinite_def HInfinite_def) |
1272 |
apply (auto dest: order_less_trans) |
|
1273 |
done |
|
1274 |
||
1275 |
lemma HFinite_not_HInfinite: |
|
1276 |
assumes x: "x \<in> HFinite" shows "x \<notin> HInfinite" |
|
1277 |
proof |
|
1278 |
assume x': "x \<in> HInfinite" |
|
1279 |
with x have "x \<in> HFinite \<inter> HInfinite" by blast |
|
1280 |
thus False by auto |
|
1281 |
qed |
|
1282 |
||
1283 |
lemma not_HFinite_HInfinite: "x\<notin> HFinite ==> x \<in> HInfinite" |
|
1284 |
apply (simp add: HInfinite_def HFinite_def, auto) |
|
1285 |
apply (drule_tac x = "r + 1" in bspec) |
|
| 15539 | 1286 |
apply (auto) |
| 14370 | 1287 |
done |
1288 |
||
1289 |
lemma HInfinite_HFinite_disj: "x \<in> HInfinite | x \<in> HFinite" |
|
1290 |
by (blast intro: not_HFinite_HInfinite) |
|
1291 |
||
1292 |
lemma HInfinite_HFinite_iff: "(x \<in> HInfinite) = (x \<notin> HFinite)" |
|
1293 |
by (blast dest: HFinite_not_HInfinite not_HFinite_HInfinite) |
|
1294 |
||
1295 |
lemma HFinite_HInfinite_iff: "(x \<in> HFinite) = (x \<notin> HInfinite)" |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1296 |
by (simp add: HInfinite_HFinite_iff) |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1297 |
|
| 14370 | 1298 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1299 |
lemma HInfinite_diff_HFinite_Infinitesimal_disj: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1300 |
"x \<notin> Infinitesimal ==> x \<in> HInfinite | x \<in> HFinite - Infinitesimal" |
| 14370 | 1301 |
by (fast intro: not_HFinite_HInfinite) |
1302 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1303 |
lemma HFinite_inverse: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1304 |
fixes x :: "'a::real_normed_div_algebra star" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1305 |
shows "[| x \<in> HFinite; x \<notin> Infinitesimal |] ==> inverse x \<in> HFinite" |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1306 |
apply (subgoal_tac "x \<noteq> 0") |
| 14370 | 1307 |
apply (cut_tac x = "inverse x" in HInfinite_HFinite_disj) |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1308 |
apply (auto dest!: HInfinite_inverse_Infinitesimal |
|
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1309 |
simp add: nonzero_inverse_inverse_eq) |
| 14370 | 1310 |
done |
1311 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1312 |
lemma HFinite_inverse2: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1313 |
fixes x :: "'a::real_normed_div_algebra star" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1314 |
shows "x \<in> HFinite - Infinitesimal ==> inverse x \<in> HFinite" |
| 14370 | 1315 |
by (blast intro: HFinite_inverse) |
1316 |
||
1317 |
(* stronger statement possible in fact *) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1318 |
lemma Infinitesimal_inverse_HFinite: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1319 |
fixes x :: "'a::real_normed_div_algebra star" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1320 |
shows "x \<notin> Infinitesimal ==> inverse(x) \<in> HFinite" |
| 14370 | 1321 |
apply (drule HInfinite_diff_HFinite_Infinitesimal_disj) |
1322 |
apply (blast intro: HFinite_inverse HInfinite_inverse_Infinitesimal Infinitesimal_subset_HFinite [THEN subsetD]) |
|
1323 |
done |
|
1324 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1325 |
lemma HFinite_not_Infinitesimal_inverse: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1326 |
fixes x :: "'a::real_normed_div_algebra star" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1327 |
shows "x \<in> HFinite - Infinitesimal ==> inverse x \<in> HFinite - Infinitesimal" |
| 14370 | 1328 |
apply (auto intro: Infinitesimal_inverse_HFinite) |
1329 |
apply (drule Infinitesimal_HFinite_mult2, assumption) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1330 |
apply (simp add: not_Infinitesimal_not_zero right_inverse) |
| 14370 | 1331 |
done |
1332 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1333 |
lemma approx_inverse: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1334 |
fixes x y :: "'a::real_normed_div_algebra star" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1335 |
shows |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1336 |
"[| x @= y; y \<in> HFinite - Infinitesimal |] |
| 14370 | 1337 |
==> inverse x @= inverse y" |
1338 |
apply (frule HFinite_diff_Infinitesimal_approx, assumption) |
|
1339 |
apply (frule not_Infinitesimal_not_zero2) |
|
1340 |
apply (frule_tac x = x in not_Infinitesimal_not_zero2) |
|
1341 |
apply (drule HFinite_inverse2)+ |
|
1342 |
apply (drule approx_mult2, assumption, auto) |
|
1343 |
apply (drule_tac c = "inverse x" in approx_mult1, assumption) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1344 |
apply (auto intro: approx_sym simp add: mult_assoc) |
| 14370 | 1345 |
done |
1346 |
||
1347 |
(*Used for NSLIM_inverse, NSLIMSEQ_inverse*) |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1348 |
lemmas star_of_approx_inverse = star_of_HFinite_diff_Infinitesimal [THEN [2] approx_inverse] |
| 14370 | 1349 |
lemmas hypreal_of_real_approx_inverse = hypreal_of_real_HFinite_diff_Infinitesimal [THEN [2] approx_inverse] |
1350 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1351 |
lemma inverse_add_Infinitesimal_approx: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1352 |
fixes x h :: "'a::real_normed_div_algebra star" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1353 |
shows |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1354 |
"[| x \<in> HFinite - Infinitesimal; |
| 14370 | 1355 |
h \<in> Infinitesimal |] ==> inverse(x + h) @= inverse x" |
1356 |
apply (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self) |
|
1357 |
done |
|
1358 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1359 |
lemma inverse_add_Infinitesimal_approx2: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1360 |
fixes x h :: "'a::real_normed_div_algebra star" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1361 |
shows |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1362 |
"[| x \<in> HFinite - Infinitesimal; |
| 14370 | 1363 |
h \<in> Infinitesimal |] ==> inverse(h + x) @= inverse x" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1364 |
apply (rule add_commute [THEN subst]) |
| 14370 | 1365 |
apply (blast intro: inverse_add_Infinitesimal_approx) |
1366 |
done |
|
1367 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1368 |
lemma inverse_add_Infinitesimal_approx_Infinitesimal: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1369 |
fixes x h :: "'a::real_normed_div_algebra star" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1370 |
shows |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1371 |
"[| x \<in> HFinite - Infinitesimal; |
| 20563 | 1372 |
h \<in> Infinitesimal |] ==> inverse(x + h) - inverse x @= h" |
| 14370 | 1373 |
apply (rule approx_trans2) |
| 15229 | 1374 |
apply (auto intro: inverse_add_Infinitesimal_approx |
1375 |
simp add: mem_infmal_iff approx_minus_iff [symmetric]) |
|
| 14370 | 1376 |
done |
1377 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1378 |
lemma Infinitesimal_square_iff: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1379 |
fixes x :: "'a::real_normed_div_algebra star" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1380 |
shows "(x \<in> Infinitesimal) = (x*x \<in> Infinitesimal)" |
| 14370 | 1381 |
apply (auto intro: Infinitesimal_mult) |
1382 |
apply (rule ccontr, frule Infinitesimal_inverse_HFinite) |
|
1383 |
apply (frule not_Infinitesimal_not_zero) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1384 |
apply (auto dest: Infinitesimal_HFinite_mult simp add: mult_assoc) |
| 14370 | 1385 |
done |
1386 |
declare Infinitesimal_square_iff [symmetric, simp] |
|
1387 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1388 |
lemma HFinite_square_iff [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1389 |
fixes x :: "'a::real_normed_div_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1390 |
shows "(x*x \<in> HFinite) = (x \<in> HFinite)" |
| 14370 | 1391 |
apply (auto intro: HFinite_mult) |
1392 |
apply (auto dest: HInfinite_mult simp add: HFinite_HInfinite_iff) |
|
1393 |
done |
|
1394 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1395 |
lemma HInfinite_square_iff [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1396 |
fixes x :: "'a::real_normed_div_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1397 |
shows "(x*x \<in> HInfinite) = (x \<in> HInfinite)" |
| 14370 | 1398 |
by (auto simp add: HInfinite_HFinite_iff) |
1399 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1400 |
lemma approx_HFinite_mult_cancel: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1401 |
fixes a w z :: "'a::real_normed_div_algebra star" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1402 |
shows "[| a: HFinite-Infinitesimal; a* w @= a*z |] ==> w @= z" |
| 14370 | 1403 |
apply safe |
1404 |
apply (frule HFinite_inverse, assumption) |
|
1405 |
apply (drule not_Infinitesimal_not_zero) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1406 |
apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) |
| 14370 | 1407 |
done |
1408 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1409 |
lemma approx_HFinite_mult_cancel_iff1: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1410 |
fixes a w z :: "'a::real_normed_div_algebra star" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1411 |
shows "a: HFinite-Infinitesimal ==> (a * w @= a * z) = (w @= z)" |
| 14370 | 1412 |
by (auto intro: approx_mult2 approx_HFinite_mult_cancel) |
1413 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1414 |
lemma HInfinite_HFinite_add_cancel: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1415 |
"[| x + y \<in> HInfinite; y \<in> HFinite |] ==> x \<in> HInfinite" |
| 14370 | 1416 |
apply (rule ccontr) |
1417 |
apply (drule HFinite_HInfinite_iff [THEN iffD2]) |
|
1418 |
apply (auto dest: HFinite_add simp add: HInfinite_HFinite_iff) |
|
1419 |
done |
|
1420 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1421 |
lemma HInfinite_HFinite_add: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1422 |
"[| x \<in> HInfinite; y \<in> HFinite |] ==> x + y \<in> HInfinite" |
| 14370 | 1423 |
apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1424 |
apply (auto simp add: add_assoc HFinite_minus_iff) |
| 14370 | 1425 |
done |
1426 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1427 |
lemma HInfinite_ge_HInfinite: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1428 |
"[| (x::hypreal) \<in> HInfinite; x \<le> y; 0 \<le> x |] ==> y \<in> HInfinite" |
| 14370 | 1429 |
by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff) |
1430 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1431 |
lemma Infinitesimal_inverse_HInfinite: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1432 |
fixes x :: "'a::real_normed_div_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1433 |
shows "[| x \<in> Infinitesimal; x \<noteq> 0 |] ==> inverse x \<in> HInfinite" |
| 14370 | 1434 |
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2]) |
1435 |
apply (auto dest: Infinitesimal_HFinite_mult2) |
|
1436 |
done |
|
1437 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1438 |
lemma HInfinite_HFinite_not_Infinitesimal_mult: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1439 |
fixes x y :: "'a::real_normed_div_algebra star" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1440 |
shows "[| x \<in> HInfinite; y \<in> HFinite - Infinitesimal |] |
| 14370 | 1441 |
==> x * y \<in> HInfinite" |
1442 |
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2]) |
|
1443 |
apply (frule HFinite_Infinitesimal_not_zero) |
|
1444 |
apply (drule HFinite_not_Infinitesimal_inverse) |
|
1445 |
apply (safe, drule HFinite_mult) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1446 |
apply (auto simp add: mult_assoc HFinite_HInfinite_iff) |
| 14370 | 1447 |
done |
1448 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1449 |
lemma HInfinite_HFinite_not_Infinitesimal_mult2: |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1450 |
fixes x y :: "'a::real_normed_div_algebra star" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1451 |
shows "[| x \<in> HInfinite; y \<in> HFinite - Infinitesimal |] |
| 14370 | 1452 |
==> y * x \<in> HInfinite" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1453 |
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2]) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1454 |
apply (frule HFinite_Infinitesimal_not_zero) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1455 |
apply (drule HFinite_not_Infinitesimal_inverse) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1456 |
apply (safe, drule_tac x="inverse y" in HFinite_mult) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1457 |
apply assumption |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1458 |
apply (auto simp add: mult_assoc [symmetric] HFinite_HInfinite_iff) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1459 |
done |
| 14370 | 1460 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1461 |
lemma HInfinite_gt_SReal: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1462 |
"[| (x::hypreal) \<in> HInfinite; 0 < x; y \<in> Reals |] ==> y < x" |
| 15003 | 1463 |
by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le) |
| 14370 | 1464 |
|
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1465 |
lemma HInfinite_gt_zero_gt_one: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1466 |
"[| (x::hypreal) \<in> HInfinite; 0 < x |] ==> 1 < x" |
| 14370 | 1467 |
by (auto intro: HInfinite_gt_SReal) |
1468 |
||
1469 |
||
| 15229 | 1470 |
lemma not_HInfinite_one [simp]: "1 \<notin> HInfinite" |
| 14370 | 1471 |
apply (simp (no_asm) add: HInfinite_HFinite_iff) |
1472 |
done |
|
1473 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1474 |
lemma approx_hrabs_disj: "abs (x::hypreal) @= x | abs x @= -x" |
| 14370 | 1475 |
by (cut_tac x = x in hrabs_disj, auto) |
1476 |
||
1477 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1478 |
subsection{*Theorems about Monads*}
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1479 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1480 |
lemma monad_hrabs_Un_subset: "monad (abs x) \<le> monad(x::hypreal) Un monad(-x)" |
| 14370 | 1481 |
by (rule_tac x1 = x in hrabs_disj [THEN disjE], auto) |
1482 |
||
1483 |
lemma Infinitesimal_monad_eq: "e \<in> Infinitesimal ==> monad (x+e) = monad x" |
|
1484 |
by (fast intro!: Infinitesimal_add_approx_self [THEN approx_sym] approx_monad_iff [THEN iffD1]) |
|
1485 |
||
1486 |
lemma mem_monad_iff: "(u \<in> monad x) = (-u \<in> monad (-x))" |
|
1487 |
by (simp add: monad_def) |
|
1488 |
||
1489 |
lemma Infinitesimal_monad_zero_iff: "(x \<in> Infinitesimal) = (x \<in> monad 0)" |
|
1490 |
by (auto intro: approx_sym simp add: monad_def mem_infmal_iff) |
|
1491 |
||
1492 |
lemma monad_zero_minus_iff: "(x \<in> monad 0) = (-x \<in> monad 0)" |
|
1493 |
apply (simp (no_asm) add: Infinitesimal_monad_zero_iff [symmetric]) |
|
1494 |
done |
|
1495 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1496 |
lemma monad_zero_hrabs_iff: "((x::hypreal) \<in> monad 0) = (abs x \<in> monad 0)" |
| 14370 | 1497 |
apply (rule_tac x1 = x in hrabs_disj [THEN disjE]) |
1498 |
apply (auto simp add: monad_zero_minus_iff [symmetric]) |
|
1499 |
done |
|
1500 |
||
| 15229 | 1501 |
lemma mem_monad_self [simp]: "x \<in> monad x" |
| 14370 | 1502 |
by (simp add: monad_def) |
| 15229 | 1503 |
|
| 14370 | 1504 |
|
| 15229 | 1505 |
subsection{*Proof that @{term "x @= y"} implies @{term"\<bar>x\<bar> @= \<bar>y\<bar>"}*}
|
1506 |
||
1507 |
lemma approx_subset_monad: "x @= y ==> {x,y} \<le> monad x"
|
|
| 14370 | 1508 |
apply (simp (no_asm)) |
1509 |
apply (simp add: approx_monad_iff) |
|
1510 |
done |
|
1511 |
||
| 15229 | 1512 |
lemma approx_subset_monad2: "x @= y ==> {x,y} \<le> monad y"
|
| 14370 | 1513 |
apply (drule approx_sym) |
1514 |
apply (fast dest: approx_subset_monad) |
|
1515 |
done |
|
1516 |
||
1517 |
lemma mem_monad_approx: "u \<in> monad x ==> x @= u" |
|
1518 |
by (simp add: monad_def) |
|
1519 |
||
1520 |
lemma approx_mem_monad: "x @= u ==> u \<in> monad x" |
|
1521 |
by (simp add: monad_def) |
|
1522 |
||
1523 |
lemma approx_mem_monad2: "x @= u ==> x \<in> monad u" |
|
1524 |
apply (simp add: monad_def) |
|
1525 |
apply (blast intro!: approx_sym) |
|
1526 |
done |
|
1527 |
||
1528 |
lemma approx_mem_monad_zero: "[| x @= y;x \<in> monad 0 |] ==> y \<in> monad 0" |
|
1529 |
apply (drule mem_monad_approx) |
|
1530 |
apply (fast intro: approx_mem_monad approx_trans) |
|
1531 |
done |
|
1532 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1533 |
lemma Infinitesimal_approx_hrabs: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1534 |
"[| x @= y; (x::hypreal) \<in> Infinitesimal |] ==> abs x @= abs y" |
| 14370 | 1535 |
apply (drule Infinitesimal_monad_zero_iff [THEN iffD1]) |
1536 |
apply (blast intro: approx_mem_monad_zero monad_zero_hrabs_iff [THEN iffD1] mem_monad_approx approx_trans3) |
|
1537 |
done |
|
1538 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1539 |
lemma less_Infinitesimal_less: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1540 |
"[| 0 < x; (x::hypreal) \<notin>Infinitesimal; e :Infinitesimal |] ==> e < x" |
| 14370 | 1541 |
apply (rule ccontr) |
1542 |
apply (auto intro: Infinitesimal_zero [THEN [2] Infinitesimal_interval] |
|
1543 |
dest!: order_le_imp_less_or_eq simp add: linorder_not_less) |
|
1544 |
done |
|
1545 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1546 |
lemma Ball_mem_monad_gt_zero: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1547 |
"[| 0 < (x::hypreal); x \<notin> Infinitesimal; u \<in> monad x |] ==> 0 < u" |
| 14370 | 1548 |
apply (drule mem_monad_approx [THEN approx_sym]) |
1549 |
apply (erule bex_Infinitesimal_iff2 [THEN iffD2, THEN bexE]) |
|
1550 |
apply (drule_tac e = "-xa" in less_Infinitesimal_less, auto) |
|
1551 |
done |
|
1552 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1553 |
lemma Ball_mem_monad_less_zero: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1554 |
"[| (x::hypreal) < 0; x \<notin> Infinitesimal; u \<in> monad x |] ==> u < 0" |
| 14370 | 1555 |
apply (drule mem_monad_approx [THEN approx_sym]) |
1556 |
apply (erule bex_Infinitesimal_iff [THEN iffD2, THEN bexE]) |
|
1557 |
apply (cut_tac x = "-x" and e = xa in less_Infinitesimal_less, auto) |
|
1558 |
done |
|
1559 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1560 |
lemma lemma_approx_gt_zero: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1561 |
"[|0 < (x::hypreal); x \<notin> Infinitesimal; x @= y|] ==> 0 < y" |
| 14370 | 1562 |
by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad) |
1563 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1564 |
lemma lemma_approx_less_zero: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1565 |
"[|(x::hypreal) < 0; x \<notin> Infinitesimal; x @= y|] ==> y < 0" |
| 14370 | 1566 |
by (blast dest: Ball_mem_monad_less_zero approx_subset_monad) |
1567 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1568 |
theorem approx_hrabs: "(x::hypreal) @= y ==> abs x @= abs y" |
|
20652
6e9b7617c89a
added approx_hnorm theorem; removed division_by_zero class requirements from several lemmas
huffman
parents:
20633
diff
changeset
|
1569 |
by (drule approx_hnorm, simp) |
| 14370 | 1570 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1571 |
lemma approx_hrabs_zero_cancel: "abs(x::hypreal) @= 0 ==> x @= 0" |
| 14370 | 1572 |
apply (cut_tac x = x in hrabs_disj) |
1573 |
apply (auto dest: approx_minus) |
|
1574 |
done |
|
1575 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1576 |
lemma approx_hrabs_add_Infinitesimal: |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1577 |
"(e::hypreal) \<in> Infinitesimal ==> abs x @= abs(x+e)" |
| 14370 | 1578 |
by (fast intro: approx_hrabs Infinitesimal_add_approx_self) |
1579 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1580 |
lemma approx_hrabs_add_minus_Infinitesimal: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1581 |
"(e::hypreal) \<in> Infinitesimal ==> abs x @= abs(x + -e)" |
| 14370 | 1582 |
by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self) |
1583 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1584 |
lemma hrabs_add_Infinitesimal_cancel: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1585 |
"[| (e::hypreal) \<in> Infinitesimal; e' \<in> Infinitesimal; |
| 14370 | 1586 |
abs(x+e) = abs(y+e')|] ==> abs x @= abs y" |
1587 |
apply (drule_tac x = x in approx_hrabs_add_Infinitesimal) |
|
1588 |
apply (drule_tac x = y in approx_hrabs_add_Infinitesimal) |
|
1589 |
apply (auto intro: approx_trans2) |
|
1590 |
done |
|
1591 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1592 |
lemma hrabs_add_minus_Infinitesimal_cancel: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1593 |
"[| (e::hypreal) \<in> Infinitesimal; e' \<in> Infinitesimal; |
| 14370 | 1594 |
abs(x + -e) = abs(y + -e')|] ==> abs x @= abs y" |
1595 |
apply (drule_tac x = x in approx_hrabs_add_minus_Infinitesimal) |
|
1596 |
apply (drule_tac x = y in approx_hrabs_add_minus_Infinitesimal) |
|
1597 |
apply (auto intro: approx_trans2) |
|
1598 |
done |
|
1599 |
||
1600 |
(* interesting slightly counterintuitive theorem: necessary |
|
1601 |
for proving that an open interval is an NS open set |
|
1602 |
*) |
|
1603 |
lemma Infinitesimal_add_hypreal_of_real_less: |
|
1604 |
"[| x < y; u \<in> Infinitesimal |] |
|
1605 |
==> hypreal_of_real x + u < hypreal_of_real y" |
|
1606 |
apply (simp add: Infinitesimal_def) |
|
| 17431 | 1607 |
apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec, simp) |
|
20254
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
webertj
parents:
20217
diff
changeset
|
1608 |
apply (simp add: abs_less_iff) |
| 14370 | 1609 |
done |
1610 |
||
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
1611 |
lemma Infinitesimal_add_hrabs_hypreal_of_real_less: |
|
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
1612 |
"[| x \<in> Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] |
| 14370 | 1613 |
==> abs (hypreal_of_real r + x) < hypreal_of_real y" |
1614 |
apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal) |
|
1615 |
apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]]) |
|
| 15229 | 1616 |
apply (auto intro!: Infinitesimal_add_hypreal_of_real_less |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1617 |
simp del: star_of_abs |
| 15229 | 1618 |
simp add: hypreal_of_real_hrabs) |
| 14370 | 1619 |
done |
1620 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1621 |
lemma Infinitesimal_add_hrabs_hypreal_of_real_less2: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1622 |
"[| x \<in> Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] |
| 14370 | 1623 |
==> abs (x + hypreal_of_real r) < hypreal_of_real y" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1624 |
apply (rule add_commute [THEN subst]) |
| 14370 | 1625 |
apply (erule Infinitesimal_add_hrabs_hypreal_of_real_less, assumption) |
1626 |
done |
|
1627 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1628 |
lemma hypreal_of_real_le_add_Infininitesimal_cancel: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1629 |
"[| u \<in> Infinitesimal; v \<in> Infinitesimal; |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1630 |
hypreal_of_real x + u \<le> hypreal_of_real y + v |] |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1631 |
==> hypreal_of_real x \<le> hypreal_of_real y" |
| 14370 | 1632 |
apply (simp add: linorder_not_less [symmetric], auto) |
1633 |
apply (drule_tac u = "v-u" in Infinitesimal_add_hypreal_of_real_less) |
|
1634 |
apply (auto simp add: Infinitesimal_diff) |
|
1635 |
done |
|
1636 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1637 |
lemma hypreal_of_real_le_add_Infininitesimal_cancel2: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1638 |
"[| u \<in> Infinitesimal; v \<in> Infinitesimal; |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1639 |
hypreal_of_real x + u \<le> hypreal_of_real y + v |] |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1640 |
==> x \<le> y" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1641 |
by (blast intro: star_of_le [THEN iffD1] |
|
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1642 |
intro!: hypreal_of_real_le_add_Infininitesimal_cancel) |
| 14370 | 1643 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1644 |
lemma hypreal_of_real_less_Infinitesimal_le_zero: |
| 15229 | 1645 |
"[| hypreal_of_real x < e; e \<in> Infinitesimal |] ==> hypreal_of_real x \<le> 0" |
| 14370 | 1646 |
apply (rule linorder_not_less [THEN iffD1], safe) |
1647 |
apply (drule Infinitesimal_interval) |
|
1648 |
apply (drule_tac [4] SReal_hypreal_of_real [THEN SReal_Infinitesimal_zero], auto) |
|
1649 |
done |
|
1650 |
||
1651 |
(*used once, in Lim/NSDERIV_inverse*) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1652 |
lemma Infinitesimal_add_not_zero: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1653 |
"[| h \<in> Infinitesimal; x \<noteq> 0 |] ==> hypreal_of_real x + h \<noteq> 0" |
| 14370 | 1654 |
apply auto |
1655 |
apply (subgoal_tac "h = - hypreal_of_real x", auto) |
|
1656 |
done |
|
1657 |
||
| 15229 | 1658 |
lemma Infinitesimal_square_cancel [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1659 |
"(x::hypreal)*x + y*y \<in> Infinitesimal ==> x*x \<in> Infinitesimal" |
| 14370 | 1660 |
apply (rule Infinitesimal_interval2) |
1661 |
apply (rule_tac [3] zero_le_square, assumption) |
|
1662 |
apply (auto simp add: zero_le_square) |
|
1663 |
done |
|
1664 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1665 |
lemma HFinite_square_cancel [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1666 |
"(x::hypreal)*x + y*y \<in> HFinite ==> x*x \<in> HFinite" |
| 14370 | 1667 |
apply (rule HFinite_bounded, assumption) |
1668 |
apply (auto simp add: zero_le_square) |
|
1669 |
done |
|
1670 |
||
| 15229 | 1671 |
lemma Infinitesimal_square_cancel2 [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1672 |
"(x::hypreal)*x + y*y \<in> Infinitesimal ==> y*y \<in> Infinitesimal" |
| 14370 | 1673 |
apply (rule Infinitesimal_square_cancel) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1674 |
apply (rule add_commute [THEN subst]) |
| 14370 | 1675 |
apply (simp (no_asm)) |
1676 |
done |
|
1677 |
||
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1678 |
lemma HFinite_square_cancel2 [simp]: |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1679 |
"(x::hypreal)*x + y*y \<in> HFinite ==> y*y \<in> HFinite" |
| 14370 | 1680 |
apply (rule HFinite_square_cancel) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1681 |
apply (rule add_commute [THEN subst]) |
| 14370 | 1682 |
apply (simp (no_asm)) |
1683 |
done |
|
1684 |
||
| 15229 | 1685 |
lemma Infinitesimal_sum_square_cancel [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1686 |
"(x::hypreal)*x + y*y + z*z \<in> Infinitesimal ==> x*x \<in> Infinitesimal" |
| 14370 | 1687 |
apply (rule Infinitesimal_interval2, assumption) |
1688 |
apply (rule_tac [2] zero_le_square, simp) |
|
1689 |
apply (insert zero_le_square [of y]) |
|
1690 |
apply (insert zero_le_square [of z], simp) |
|
1691 |
done |
|
1692 |
||
| 15229 | 1693 |
lemma HFinite_sum_square_cancel [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1694 |
"(x::hypreal)*x + y*y + z*z \<in> HFinite ==> x*x \<in> HFinite" |
| 14370 | 1695 |
apply (rule HFinite_bounded, assumption) |
1696 |
apply (rule_tac [2] zero_le_square) |
|
1697 |
apply (insert zero_le_square [of y]) |
|
1698 |
apply (insert zero_le_square [of z], simp) |
|
1699 |
done |
|
1700 |
||
| 15229 | 1701 |
lemma Infinitesimal_sum_square_cancel2 [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1702 |
"(y::hypreal)*y + x*x + z*z \<in> Infinitesimal ==> x*x \<in> Infinitesimal" |
| 14370 | 1703 |
apply (rule Infinitesimal_sum_square_cancel) |
1704 |
apply (simp add: add_ac) |
|
1705 |
done |
|
1706 |
||
| 15229 | 1707 |
lemma HFinite_sum_square_cancel2 [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1708 |
"(y::hypreal)*y + x*x + z*z \<in> HFinite ==> x*x \<in> HFinite" |
| 14370 | 1709 |
apply (rule HFinite_sum_square_cancel) |
1710 |
apply (simp add: add_ac) |
|
1711 |
done |
|
1712 |
||
| 15229 | 1713 |
lemma Infinitesimal_sum_square_cancel3 [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1714 |
"(z::hypreal)*z + y*y + x*x \<in> Infinitesimal ==> x*x \<in> Infinitesimal" |
| 14370 | 1715 |
apply (rule Infinitesimal_sum_square_cancel) |
1716 |
apply (simp add: add_ac) |
|
1717 |
done |
|
1718 |
||
| 15229 | 1719 |
lemma HFinite_sum_square_cancel3 [simp]: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1720 |
"(z::hypreal)*z + y*y + x*x \<in> HFinite ==> x*x \<in> HFinite" |
| 14370 | 1721 |
apply (rule HFinite_sum_square_cancel) |
1722 |
apply (simp add: add_ac) |
|
1723 |
done |
|
1724 |
||
| 15229 | 1725 |
lemma monad_hrabs_less: |
1726 |
"[| y \<in> monad x; 0 < hypreal_of_real e |] |
|
| 20563 | 1727 |
==> abs (y - x) < hypreal_of_real e" |
| 14370 | 1728 |
apply (drule mem_monad_approx [THEN approx_sym]) |
1729 |
apply (drule bex_Infinitesimal_iff [THEN iffD2]) |
|
1730 |
apply (auto dest!: InfinitesimalD) |
|
1731 |
done |
|
1732 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1733 |
lemma mem_monad_SReal_HFinite: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1734 |
"x \<in> monad (hypreal_of_real a) ==> x \<in> HFinite" |
| 14370 | 1735 |
apply (drule mem_monad_approx [THEN approx_sym]) |
1736 |
apply (drule bex_Infinitesimal_iff2 [THEN iffD2]) |
|
1737 |
apply (safe dest!: Infinitesimal_subset_HFinite [THEN subsetD]) |
|
1738 |
apply (erule SReal_hypreal_of_real [THEN SReal_subset_HFinite [THEN subsetD], THEN HFinite_add]) |
|
1739 |
done |
|
1740 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1741 |
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1742 |
subsection{* Theorems about Standard Part*}
|
| 14370 | 1743 |
|
1744 |
lemma st_approx_self: "x \<in> HFinite ==> st x @= x" |
|
1745 |
apply (simp add: st_def) |
|
1746 |
apply (frule st_part_Ex, safe) |
|
1747 |
apply (rule someI2) |
|
1748 |
apply (auto intro: approx_sym) |
|
1749 |
done |
|
1750 |
||
1751 |
lemma st_SReal: "x \<in> HFinite ==> st x \<in> Reals" |
|
1752 |
apply (simp add: st_def) |
|
1753 |
apply (frule st_part_Ex, safe) |
|
1754 |
apply (rule someI2) |
|
1755 |
apply (auto intro: approx_sym) |
|
1756 |
done |
|
1757 |
||
1758 |
lemma st_HFinite: "x \<in> HFinite ==> st x \<in> HFinite" |
|
1759 |
by (erule st_SReal [THEN SReal_subset_HFinite [THEN subsetD]]) |
|
1760 |
||
1761 |
lemma st_SReal_eq: "x \<in> Reals ==> st x = x" |
|
1762 |
apply (simp add: st_def) |
|
1763 |
apply (rule some_equality) |
|
1764 |
apply (fast intro: SReal_subset_HFinite [THEN subsetD]) |
|
1765 |
apply (blast dest: SReal_approx_iff [THEN iffD1]) |
|
1766 |
done |
|
1767 |
||
| 15229 | 1768 |
lemma st_hypreal_of_real [simp]: "st (hypreal_of_real x) = hypreal_of_real x" |
| 14370 | 1769 |
by (rule SReal_hypreal_of_real [THEN st_SReal_eq]) |
1770 |
||
1771 |
lemma st_eq_approx: "[| x \<in> HFinite; y \<in> HFinite; st x = st y |] ==> x @= y" |
|
1772 |
by (auto dest!: st_approx_self elim!: approx_trans3) |
|
1773 |
||
1774 |
lemma approx_st_eq: |
|
1775 |
assumes "x \<in> HFinite" and "y \<in> HFinite" and "x @= y" |
|
1776 |
shows "st x = st y" |
|
1777 |
proof - |
|
1778 |
have "st x @= x" "st y @= y" "st x \<in> Reals" "st y \<in> Reals" |
|
1779 |
by (simp_all add: st_approx_self st_SReal prems) |
|
1780 |
with prems show ?thesis |
|
1781 |
by (fast elim: approx_trans approx_trans2 SReal_approx_iff [THEN iffD1]) |
|
1782 |
qed |
|
1783 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1784 |
lemma st_eq_approx_iff: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1785 |
"[| x \<in> HFinite; y \<in> HFinite|] |
| 14370 | 1786 |
==> (x @= y) = (st x = st y)" |
1787 |
by (blast intro: approx_st_eq st_eq_approx) |
|
1788 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1789 |
lemma st_Infinitesimal_add_SReal: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1790 |
"[| x \<in> Reals; e \<in> Infinitesimal |] ==> st(x + e) = x" |
| 14370 | 1791 |
apply (frule st_SReal_eq [THEN subst]) |
1792 |
prefer 2 apply assumption |
|
1793 |
apply (frule SReal_subset_HFinite [THEN subsetD]) |
|
1794 |
apply (frule Infinitesimal_subset_HFinite [THEN subsetD]) |
|
1795 |
apply (drule st_SReal_eq) |
|
1796 |
apply (rule approx_st_eq) |
|
1797 |
apply (auto intro: HFinite_add simp add: Infinitesimal_add_approx_self [THEN approx_sym]) |
|
1798 |
done |
|
1799 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1800 |
lemma st_Infinitesimal_add_SReal2: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1801 |
"[| x \<in> Reals; e \<in> Infinitesimal |] ==> st(e + x) = x" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1802 |
apply (rule add_commute [THEN subst]) |
| 14370 | 1803 |
apply (blast intro!: st_Infinitesimal_add_SReal) |
1804 |
done |
|
1805 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1806 |
lemma HFinite_st_Infinitesimal_add: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1807 |
"x \<in> HFinite ==> \<exists>e \<in> Infinitesimal. x = st(x) + e" |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1808 |
by (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2]) |
| 14370 | 1809 |
|
1810 |
lemma st_add: |
|
1811 |
assumes x: "x \<in> HFinite" and y: "y \<in> HFinite" |
|
1812 |
shows "st (x + y) = st(x) + st(y)" |
|
1813 |
proof - |
|
1814 |
from HFinite_st_Infinitesimal_add [OF x] |
|
1815 |
obtain ex where ex: "ex \<in> Infinitesimal" "st x + ex = x" |
|
1816 |
by (blast intro: sym) |
|
1817 |
from HFinite_st_Infinitesimal_add [OF y] |
|
1818 |
obtain ey where ey: "ey \<in> Infinitesimal" "st y + ey = y" |
|
1819 |
by (blast intro: sym) |
|
1820 |
have "st (x + y) = st ((st x + ex) + (st y + ey))" |
|
1821 |
by (simp add: ex ey) |
|
1822 |
also have "... = st ((ex + ey) + (st x + st y))" by (simp add: add_ac) |
|
1823 |
also have "... = st x + st y" |
|
| 15539 | 1824 |
by (simp add: prems st_SReal Infinitesimal_add |
| 14370 | 1825 |
st_Infinitesimal_add_SReal2) |
1826 |
finally show ?thesis . |
|
1827 |
qed |
|
1828 |
||
| 15229 | 1829 |
lemma st_number_of [simp]: "st (number_of w) = number_of w" |
| 14370 | 1830 |
by (rule SReal_number_of [THEN st_SReal_eq]) |
1831 |
||
1832 |
(*the theorem above for the special cases of zero and one*) |
|
1833 |
lemma [simp]: "st 0 = 0" "st 1 = 1" |
|
1834 |
by (simp_all add: st_SReal_eq) |
|
1835 |
||
1836 |
lemma st_minus: assumes "y \<in> HFinite" shows "st(-y) = -st(y)" |
|
1837 |
proof - |
|
1838 |
have "st (- y) + st y = 0" |
|
1839 |
by (simp add: prems st_add [symmetric] HFinite_minus_iff) |
|
1840 |
thus ?thesis by arith |
|
1841 |
qed |
|
1842 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1843 |
lemma st_diff: "[| x \<in> HFinite; y \<in> HFinite |] ==> st (x-y) = st(x) - st(y)" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1844 |
apply (simp add: diff_def) |
| 14370 | 1845 |
apply (frule_tac y1 = y in st_minus [symmetric]) |
1846 |
apply (drule_tac x1 = y in HFinite_minus_iff [THEN iffD2]) |
|
1847 |
apply (simp (no_asm_simp) add: st_add) |
|
1848 |
done |
|
1849 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1850 |
lemma lemma_st_mult: |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1851 |
fixes x y e ea :: "'a::real_normed_algebra star" |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1852 |
shows "[| x \<in> HFinite; y \<in> HFinite; e \<in> Infinitesimal; ea \<in> Infinitesimal |] |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1853 |
==> e*y + x*ea + e*ea \<in> Infinitesimal" |
|
20541
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1854 |
apply (intro Infinitesimal_add) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1855 |
apply (erule (1) Infinitesimal_HFinite_mult) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1856 |
apply (erule (1) Infinitesimal_HFinite_mult2) |
|
f614c619b1e1
generalized types of Infinitesimal, HFinite, and HInfinite to work over nonstandard extensions of any real normed vector space
huffman
parents:
20485
diff
changeset
|
1857 |
apply (erule (1) Infinitesimal_mult) |
| 14370 | 1858 |
done |
1859 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1860 |
lemma st_mult: "[| x \<in> HFinite; y \<in> HFinite |] ==> st (x * y) = st(x) * st(y)" |
| 14370 | 1861 |
apply (frule HFinite_st_Infinitesimal_add) |
1862 |
apply (frule_tac x = y in HFinite_st_Infinitesimal_add, safe) |
|
1863 |
apply (subgoal_tac "st (x * y) = st ((st x + e) * (st y + ea))") |
|
1864 |
apply (drule_tac [2] sym, drule_tac [2] sym) |
|
1865 |
prefer 2 apply simp |
|
1866 |
apply (erule_tac V = "x = st x + e" in thin_rl) |
|
1867 |
apply (erule_tac V = "y = st y + ea" in thin_rl) |
|
1868 |
apply (simp add: left_distrib right_distrib) |
|
1869 |
apply (drule st_SReal)+ |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1870 |
apply (simp (no_asm_use) add: add_assoc) |
| 14370 | 1871 |
apply (rule st_Infinitesimal_add_SReal) |
1872 |
apply (blast intro!: SReal_mult) |
|
1873 |
apply (drule SReal_subset_HFinite [THEN subsetD])+ |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1874 |
apply (rule add_assoc [THEN subst]) |
| 14370 | 1875 |
apply (blast intro!: lemma_st_mult) |
1876 |
done |
|
1877 |
||
1878 |
lemma st_Infinitesimal: "x \<in> Infinitesimal ==> st x = 0" |
|
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
1879 |
apply (subst numeral_0_eq_0 [symmetric]) |
| 14370 | 1880 |
apply (rule st_number_of [THEN subst]) |
1881 |
apply (rule approx_st_eq) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1882 |
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1883 |
simp add: mem_infmal_iff [symmetric]) |
| 14370 | 1884 |
done |
1885 |
||
1886 |
lemma st_not_Infinitesimal: "st(x) \<noteq> 0 ==> x \<notin> Infinitesimal" |
|
1887 |
by (fast intro: st_Infinitesimal) |
|
1888 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1889 |
lemma st_inverse: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1890 |
"[| x \<in> HFinite; st x \<noteq> 0 |] |
| 14370 | 1891 |
==> st(inverse x) = inverse (st x)" |
1892 |
apply (rule_tac c1 = "st x" in hypreal_mult_left_cancel [THEN iffD1]) |
|
1893 |
apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1894 |
apply (subst right_inverse, auto) |
| 14370 | 1895 |
done |
1896 |
||
| 15229 | 1897 |
lemma st_divide [simp]: |
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1898 |
"[| x \<in> HFinite; y \<in> HFinite; st y \<noteq> 0 |] |
| 14370 | 1899 |
==> st(x/y) = (st x) / (st y)" |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1900 |
by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse) |
| 14370 | 1901 |
|
| 15229 | 1902 |
lemma st_idempotent [simp]: "x \<in> HFinite ==> st(st(x)) = st(x)" |
| 14370 | 1903 |
by (blast intro: st_HFinite st_approx_self approx_st_eq) |
1904 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1905 |
lemma Infinitesimal_add_st_less: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1906 |
"[| x \<in> HFinite; y \<in> HFinite; u \<in> Infinitesimal; st x < st y |] |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1907 |
==> st x + u < st y" |
| 14370 | 1908 |
apply (drule st_SReal)+ |
1909 |
apply (auto intro!: Infinitesimal_add_hypreal_of_real_less simp add: SReal_iff) |
|
1910 |
done |
|
1911 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1912 |
lemma Infinitesimal_add_st_le_cancel: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1913 |
"[| x \<in> HFinite; y \<in> HFinite; |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1914 |
u \<in> Infinitesimal; st x \<le> st y + u |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1915 |
|] ==> st x \<le> st y" |
| 14370 | 1916 |
apply (simp add: linorder_not_less [symmetric]) |
1917 |
apply (auto dest: Infinitesimal_add_st_less) |
|
1918 |
done |
|
1919 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1920 |
lemma st_le: "[| x \<in> HFinite; y \<in> HFinite; x \<le> y |] ==> st(x) \<le> st(y)" |
| 14370 | 1921 |
apply (frule HFinite_st_Infinitesimal_add) |
1922 |
apply (rotate_tac 1) |
|
1923 |
apply (frule HFinite_st_Infinitesimal_add, safe) |
|
1924 |
apply (rule Infinitesimal_add_st_le_cancel) |
|
1925 |
apply (rule_tac [3] x = ea and y = e in Infinitesimal_diff) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1926 |
apply (auto simp add: add_assoc [symmetric]) |
| 14370 | 1927 |
done |
1928 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1929 |
lemma st_zero_le: "[| 0 \<le> x; x \<in> HFinite |] ==> 0 \<le> st x" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
1930 |
apply (subst numeral_0_eq_0 [symmetric]) |
| 14370 | 1931 |
apply (rule st_number_of [THEN subst]) |
1932 |
apply (rule st_le, auto) |
|
1933 |
done |
|
1934 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1935 |
lemma st_zero_ge: "[| x \<le> 0; x \<in> HFinite |] ==> st x \<le> 0" |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14378
diff
changeset
|
1936 |
apply (subst numeral_0_eq_0 [symmetric]) |
| 14370 | 1937 |
apply (rule st_number_of [THEN subst]) |
1938 |
apply (rule st_le, auto) |
|
1939 |
done |
|
1940 |
||
1941 |
lemma st_hrabs: "x \<in> HFinite ==> abs(st x) = st(abs x)" |
|
1942 |
apply (simp add: linorder_not_le st_zero_le abs_if st_minus |
|
1943 |
linorder_not_less) |
|
1944 |
apply (auto dest!: st_zero_ge [OF order_less_imp_le]) |
|
1945 |
done |
|
1946 |
||
1947 |
||
1948 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1949 |
subsection{*Alternative Definitions for @{term HFinite} using Free Ultrafilter*}
|
| 14370 | 1950 |
|
1951 |
lemma HFinite_FreeUltrafilterNat: |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1952 |
"star_n X \<in> HFinite |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1953 |
==> \<exists>u. {n. norm (X n) < u} \<in> FreeUltrafilterNat"
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1954 |
apply (auto simp add: HFinite_def SReal_def) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1955 |
apply (rule_tac x=r in exI) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1956 |
apply (simp add: hnorm_def star_of_def starfun_star_n) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1957 |
apply (simp add: star_less_def starP2_star_n) |
| 14370 | 1958 |
done |
1959 |
||
1960 |
lemma FreeUltrafilterNat_HFinite: |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1961 |
"\<exists>u. {n. norm (X n) < u} \<in> FreeUltrafilterNat
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1962 |
==> star_n X \<in> HFinite" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1963 |
apply (auto simp add: HFinite_def mem_Rep_star_iff) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1964 |
apply (rule_tac x="star_of u" in bexI) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1965 |
apply (simp add: hnorm_def starfun_star_n star_of_def) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1966 |
apply (simp add: star_less_def starP2_star_n) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1967 |
apply (simp add: SReal_def) |
| 14370 | 1968 |
done |
1969 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1970 |
lemma HFinite_FreeUltrafilterNat_iff: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1971 |
"(star_n X \<in> HFinite) = (\<exists>u. {n. norm (X n) < u} \<in> FreeUltrafilterNat)"
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1972 |
by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite) |
| 14370 | 1973 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1974 |
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1975 |
subsection{*Alternative Definitions for @{term HInfinite} using Free Ultrafilter*}
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1976 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1977 |
lemma lemma_Compl_eq: "- {n. u < norm (xa n)} = {n. norm (xa n) \<le> u}"
|
| 14370 | 1978 |
by auto |
1979 |
||
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1980 |
lemma lemma_Compl_eq2: "- {n. norm (xa n) < u} = {n. u \<le> norm (xa n)}"
|
| 14370 | 1981 |
by auto |
1982 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1983 |
lemma lemma_Int_eq1: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1984 |
"{n. norm (xa n) \<le> u} Int {n. u \<le> norm (xa n)}
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1985 |
= {n. norm(xa n) = u}"
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
1986 |
by auto |
| 14370 | 1987 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1988 |
lemma lemma_FreeUltrafilterNat_one: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1989 |
"{n. norm (xa n) = u} \<le> {n. norm (xa n) < u + (1::real)}"
|
| 14370 | 1990 |
by auto |
1991 |
||
1992 |
(*------------------------------------- |
|
1993 |
Exclude this type of sets from free |
|
1994 |
ultrafilter for Infinite numbers! |
|
1995 |
-------------------------------------*) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
1996 |
lemma FreeUltrafilterNat_const_Finite: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
1997 |
"{n. norm (X n) = u} \<in> FreeUltrafilterNat ==> star_n X \<in> HFinite"
|
| 14370 | 1998 |
apply (rule FreeUltrafilterNat_HFinite) |
1999 |
apply (rule_tac x = "u + 1" in exI) |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2000 |
apply (erule ultra, simp) |
| 14370 | 2001 |
done |
2002 |
||
2003 |
lemma HInfinite_FreeUltrafilterNat: |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2004 |
"star_n X \<in> HInfinite ==> \<forall>u. {n. u < norm (X n)} \<in> FreeUltrafilterNat"
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2005 |
apply (drule HInfinite_HFinite_iff [THEN iffD1]) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2006 |
apply (simp add: HFinite_FreeUltrafilterNat_iff) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2007 |
apply (rule allI, drule_tac x="u + 1" in spec) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2008 |
apply (drule FreeUltrafilterNat_Compl_mem) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2009 |
apply (simp add: Collect_neg_eq [symmetric] linorder_not_less) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2010 |
apply (erule ultra, simp) |
| 14370 | 2011 |
done |
2012 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2013 |
lemma lemma_Int_HI: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2014 |
"{n. norm (Xa n) < u} Int {n. X n = Xa n} \<subseteq> {n. norm (X n) < (u::real)}"
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2015 |
by auto |
| 14370 | 2016 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2017 |
lemma lemma_Int_HIa: "{n. u < norm (X n)} Int {n. norm (X n) < u} = {}"
|
| 14370 | 2018 |
by (auto intro: order_less_asym) |
2019 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2020 |
lemma FreeUltrafilterNat_HInfinite: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2021 |
"\<forall>u. {n. u < norm (X n)} \<in> FreeUltrafilterNat ==> star_n X \<in> HInfinite"
|
| 14370 | 2022 |
apply (rule HInfinite_HFinite_iff [THEN iffD2]) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2023 |
apply (safe, drule HFinite_FreeUltrafilterNat, safe) |
| 14370 | 2024 |
apply (drule_tac x = u in spec) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2025 |
apply ultra |
| 14370 | 2026 |
done |
2027 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2028 |
lemma HInfinite_FreeUltrafilterNat_iff: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2029 |
"(star_n X \<in> HInfinite) = (\<forall>u. {n. u < norm (X n)} \<in> FreeUltrafilterNat)"
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2030 |
by (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite) |
| 14370 | 2031 |
|
2032 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2033 |
subsection{*Alternative Definitions for @{term Infinitesimal} using Free Ultrafilter*}
|
| 10751 | 2034 |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2035 |
lemma ball_SReal_eq: "(\<forall>x::hypreal \<in> Reals. P x) = (\<forall>x::real. P (star_of x))" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2036 |
by (unfold SReal_def, auto) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2037 |
|
| 14370 | 2038 |
lemma Infinitesimal_FreeUltrafilterNat: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2039 |
"star_n X \<in> Infinitesimal ==> \<forall>u>0. {n. norm (X n) < u} \<in> \<U>"
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2040 |
apply (simp add: Infinitesimal_def ball_SReal_eq) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2041 |
apply (simp add: hnorm_def starfun_star_n star_of_def) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2042 |
apply (simp add: star_less_def starP2_star_n) |
| 14370 | 2043 |
done |
2044 |
||
2045 |
lemma FreeUltrafilterNat_Infinitesimal: |
|
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2046 |
"\<forall>u>0. {n. norm (X n) < u} \<in> \<U> ==> star_n X \<in> Infinitesimal"
|
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2047 |
apply (simp add: Infinitesimal_def ball_SReal_eq) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2048 |
apply (simp add: hnorm_def starfun_star_n star_of_def) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2049 |
apply (simp add: star_less_def starP2_star_n) |
| 14370 | 2050 |
done |
2051 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2052 |
lemma Infinitesimal_FreeUltrafilterNat_iff: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2053 |
"(star_n X \<in> Infinitesimal) = (\<forall>u>0. {n. norm (X n) < u} \<in> \<U>)"
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
2054 |
by (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal) |
| 14370 | 2055 |
|
2056 |
(*------------------------------------------------------------------------ |
|
2057 |
Infinitesimals as smaller than 1/n for all n::nat (> 0) |
|
2058 |
------------------------------------------------------------------------*) |
|
2059 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2060 |
lemma lemma_Infinitesimal: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2061 |
"(\<forall>r. 0 < r --> x < r) = (\<forall>n. x < inverse(real (Suc n)))" |
| 14370 | 2062 |
apply (auto simp add: real_of_nat_Suc_gt_zero) |
2063 |
apply (blast dest!: reals_Archimedean intro: order_less_trans) |
|
2064 |
done |
|
2065 |
||
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
2066 |
lemma of_nat_in_Reals [simp]: "(of_nat n::hypreal) \<in> \<real>" |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
2067 |
apply (induct n) |
| 15539 | 2068 |
apply (simp_all) |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
2069 |
done |
|
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
2070 |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2071 |
lemma lemma_Infinitesimal2: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2072 |
"(\<forall>r \<in> Reals. 0 < r --> x < r) = |
| 14370 | 2073 |
(\<forall>n. x < inverse(hypreal_of_nat (Suc n)))" |
2074 |
apply safe |
|
2075 |
apply (drule_tac x = "inverse (hypreal_of_real (real (Suc n))) " in bspec) |
|
2076 |
apply (simp (no_asm_use) add: SReal_inverse) |
|
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
2077 |
apply (rule real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive, THEN star_of_less [THEN iffD2], THEN [2] impE]) |
| 14370 | 2078 |
prefer 2 apply assumption |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
2079 |
apply (simp add: real_of_nat_Suc_gt_zero hypreal_of_nat_eq) |
| 14370 | 2080 |
apply (auto dest!: reals_Archimedean simp add: SReal_iff) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17298
diff
changeset
|
2081 |
apply (drule star_of_less [THEN iffD2]) |
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
2082 |
apply (simp add: real_of_nat_Suc_gt_zero hypreal_of_nat_eq) |
| 14370 | 2083 |
apply (blast intro: order_less_trans) |
2084 |
done |
|
2085 |
||
|
14378
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
paulson
parents:
14371
diff
changeset
|
2086 |
|
| 14370 | 2087 |
lemma Infinitesimal_hypreal_of_nat_iff: |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2088 |
"Infinitesimal = {x. \<forall>n. hnorm x < inverse (hypreal_of_nat (Suc n))}"
|
| 14370 | 2089 |
apply (simp add: Infinitesimal_def) |
2090 |
apply (auto simp add: lemma_Infinitesimal2) |
|
2091 |
done |
|
2092 |
||
2093 |
||
| 15229 | 2094 |
subsection{*Proof that @{term omega} is an infinite number*}
|
2095 |
||
2096 |
text{*It will follow that epsilon is an infinitesimal number.*}
|
|
2097 |
||
| 14370 | 2098 |
lemma Suc_Un_eq: "{n. n < Suc m} = {n. n < m} Un {n. n = m}"
|
2099 |
by (auto simp add: less_Suc_eq) |
|
2100 |
||
2101 |
(*------------------------------------------- |
|
2102 |
Prove that any segment is finite and |
|
2103 |
hence cannot belong to FreeUltrafilterNat |
|
2104 |
-------------------------------------------*) |
|
2105 |
lemma finite_nat_segment: "finite {n::nat. n < m}"
|
|
| 15251 | 2106 |
apply (induct "m") |
| 14370 | 2107 |
apply (auto simp add: Suc_Un_eq) |
2108 |
done |
|
2109 |
||
2110 |
lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}"
|
|
2111 |
by (auto intro: finite_nat_segment) |
|
2112 |
||
2113 |
lemma finite_real_of_nat_less_real: "finite {n::nat. real n < u}"
|
|
2114 |
apply (cut_tac x = u in reals_Archimedean2, safe) |
|
2115 |
apply (rule finite_real_of_nat_segment [THEN [2] finite_subset]) |
|
2116 |
apply (auto dest: order_less_trans) |
|
2117 |
done |
|
2118 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2119 |
lemma lemma_real_le_Un_eq: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2120 |
"{n. f n \<le> u} = {n. f n < u} Un {n. u = (f n :: real)}"
|
| 14370 | 2121 |
by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) |
2122 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2123 |
lemma finite_real_of_nat_le_real: "finite {n::nat. real n \<le> u}"
|
| 14370 | 2124 |
by (auto simp add: lemma_real_le_Un_eq lemma_finite_omega_set finite_real_of_nat_less_real) |
2125 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2126 |
lemma finite_rabs_real_of_nat_le_real: "finite {n::nat. abs(real n) \<le> u}"
|
| 14370 | 2127 |
apply (simp (no_asm) add: real_of_nat_Suc_gt_zero finite_real_of_nat_le_real) |
2128 |
done |
|
2129 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2130 |
lemma rabs_real_of_nat_le_real_FreeUltrafilterNat: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2131 |
"{n. abs(real n) \<le> u} \<notin> FreeUltrafilterNat"
|
| 14370 | 2132 |
by (blast intro!: FreeUltrafilterNat_finite finite_rabs_real_of_nat_le_real) |
2133 |
||
2134 |
lemma FreeUltrafilterNat_nat_gt_real: "{n. u < real n} \<in> FreeUltrafilterNat"
|
|
2135 |
apply (rule ccontr, drule FreeUltrafilterNat_Compl_mem) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2136 |
apply (subgoal_tac "- {n::nat. u < real n} = {n. real n \<le> u}")
|
| 14370 | 2137 |
prefer 2 apply force |
2138 |
apply (simp add: finite_real_of_nat_le_real [THEN FreeUltrafilterNat_finite]) |
|
2139 |
done |
|
2140 |
||
2141 |
(*-------------------------------------------------------------- |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2142 |
The complement of {n. abs(real n) \<le> u} =
|
| 14370 | 2143 |
{n. u < abs (real n)} is in FreeUltrafilterNat
|
2144 |
by property of (free) ultrafilters |
|
2145 |
--------------------------------------------------------------*) |
|
2146 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2147 |
lemma Compl_real_le_eq: "- {n::nat. real n \<le> u} = {n. u < real n}"
|
| 14370 | 2148 |
by (auto dest!: order_le_less_trans simp add: linorder_not_le) |
2149 |
||
| 15229 | 2150 |
text{*@{term omega} is a member of @{term HInfinite}*}
|
| 14370 | 2151 |
|
2152 |
lemma FreeUltrafilterNat_omega: "{n. u < real n} \<in> FreeUltrafilterNat"
|
|
2153 |
apply (cut_tac u = u in rabs_real_of_nat_le_real_FreeUltrafilterNat) |
|
2154 |
apply (auto dest: FreeUltrafilterNat_Compl_mem simp add: Compl_real_le_eq) |
|
2155 |
done |
|
2156 |
||
| 15229 | 2157 |
theorem HInfinite_omega [simp]: "omega \<in> HInfinite" |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2158 |
apply (simp add: omega_def) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2159 |
apply (rule FreeUltrafilterNat_HInfinite) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2160 |
apply (simp (no_asm) add: real_norm_def real_of_nat_Suc diff_less_eq [symmetric] FreeUltrafilterNat_omega) |
| 14370 | 2161 |
done |
2162 |
||
2163 |
(*----------------------------------------------- |
|
2164 |
Epsilon is a member of Infinitesimal |
|
2165 |
-----------------------------------------------*) |
|
2166 |
||
| 15229 | 2167 |
lemma Infinitesimal_epsilon [simp]: "epsilon \<in> Infinitesimal" |
| 14370 | 2168 |
by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega simp add: hypreal_epsilon_inverse_omega) |
2169 |
||
| 15229 | 2170 |
lemma HFinite_epsilon [simp]: "epsilon \<in> HFinite" |
| 14370 | 2171 |
by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]) |
2172 |
||
| 15229 | 2173 |
lemma epsilon_approx_zero [simp]: "epsilon @= 0" |
| 14370 | 2174 |
apply (simp (no_asm) add: mem_infmal_iff [symmetric]) |
2175 |
done |
|
2176 |
||
2177 |
(*------------------------------------------------------------------------ |
|
2178 |
Needed for proof that we define a hyperreal [<X(n)] @= hypreal_of_real a given |
|
2179 |
that \<forall>n. |X n - a| < 1/n. Used in proof of NSLIM => LIM. |
|
2180 |
-----------------------------------------------------------------------*) |
|
2181 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2182 |
lemma real_of_nat_less_inverse_iff: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2183 |
"0 < u ==> (u < inverse (real(Suc n))) = (real(Suc n) < inverse u)" |
| 14370 | 2184 |
apply (simp add: inverse_eq_divide) |
2185 |
apply (subst pos_less_divide_eq, assumption) |
|
2186 |
apply (subst pos_less_divide_eq) |
|
2187 |
apply (simp add: real_of_nat_Suc_gt_zero) |
|
2188 |
apply (simp add: real_mult_commute) |
|
2189 |
done |
|
2190 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2191 |
lemma finite_inverse_real_of_posnat_gt_real: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2192 |
"0 < u ==> finite {n. u < inverse(real(Suc n))}"
|
| 14370 | 2193 |
apply (simp (no_asm_simp) add: real_of_nat_less_inverse_iff) |
2194 |
apply (simp (no_asm_simp) add: real_of_nat_Suc less_diff_eq [symmetric]) |
|
2195 |
apply (rule finite_real_of_nat_less_real) |
|
2196 |
done |
|
2197 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2198 |
lemma lemma_real_le_Un_eq2: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2199 |
"{n. u \<le> inverse(real(Suc n))} =
|
| 14370 | 2200 |
{n. u < inverse(real(Suc n))} Un {n. u = inverse(real(Suc n))}"
|
2201 |
apply (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) |
|
2202 |
done |
|
2203 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2204 |
lemma real_of_nat_inverse_le_iff: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2205 |
"(inverse (real(Suc n)) \<le> r) = (1 \<le> r * real(Suc n))" |
| 14370 | 2206 |
apply (simp (no_asm) add: linorder_not_less [symmetric]) |
2207 |
apply (simp (no_asm) add: inverse_eq_divide) |
|
2208 |
apply (subst pos_less_divide_eq) |
|
2209 |
apply (simp (no_asm) add: real_of_nat_Suc_gt_zero) |
|
2210 |
apply (simp (no_asm) add: real_mult_commute) |
|
2211 |
done |
|
2212 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2213 |
lemma real_of_nat_inverse_eq_iff: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2214 |
"(u = inverse (real(Suc n))) = (real(Suc n) = inverse u)" |
| 15539 | 2215 |
by (auto simp add: real_of_nat_Suc_gt_zero real_not_refl2 [THEN not_sym]) |
| 14370 | 2216 |
|
2217 |
lemma lemma_finite_omega_set2: "finite {n::nat. u = inverse(real(Suc n))}"
|
|
2218 |
apply (simp (no_asm_simp) add: real_of_nat_inverse_eq_iff) |
|
2219 |
apply (cut_tac x = "inverse u - 1" in lemma_finite_omega_set) |
|
2220 |
apply (simp add: real_of_nat_Suc diff_eq_eq [symmetric] eq_commute) |
|
2221 |
done |
|
2222 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2223 |
lemma finite_inverse_real_of_posnat_ge_real: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2224 |
"0 < u ==> finite {n. u \<le> inverse(real(Suc n))}"
|
| 14370 | 2225 |
by (auto simp add: lemma_real_le_Un_eq2 lemma_finite_omega_set2 finite_inverse_real_of_posnat_gt_real) |
2226 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2227 |
lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2228 |
"0 < u ==> {n. u \<le> inverse(real(Suc n))} \<notin> FreeUltrafilterNat"
|
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2229 |
by (blast intro!: FreeUltrafilterNat_finite finite_inverse_real_of_posnat_ge_real) |
| 14370 | 2230 |
|
2231 |
(*-------------------------------------------------------------- |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2232 |
The complement of {n. u \<le> inverse(real(Suc n))} =
|
| 14370 | 2233 |
{n. inverse(real(Suc n)) < u} is in FreeUltrafilterNat
|
2234 |
by property of (free) ultrafilters |
|
2235 |
--------------------------------------------------------------*) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2236 |
lemma Compl_le_inverse_eq: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2237 |
"- {n. u \<le> inverse(real(Suc n))} =
|
| 14370 | 2238 |
{n. inverse(real(Suc n)) < u}"
|
2239 |
apply (auto dest!: order_le_less_trans simp add: linorder_not_le) |
|
2240 |
done |
|
2241 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2242 |
lemma FreeUltrafilterNat_inverse_real_of_posnat: |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2243 |
"0 < u ==> |
| 14370 | 2244 |
{n. inverse(real(Suc n)) < u} \<in> FreeUltrafilterNat"
|
2245 |
apply (cut_tac u = u in inverse_real_of_posnat_ge_real_FreeUltrafilterNat) |
|
2246 |
apply (auto dest: FreeUltrafilterNat_Compl_mem simp add: Compl_le_inverse_eq) |
|
2247 |
done |
|
2248 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2249 |
text{* Example where we get a hyperreal from a real sequence
|
| 14370 | 2250 |
for which a particular property holds. The theorem is |
2251 |
used in proofs about equivalence of nonstandard and |
|
2252 |
standard neighbourhoods. Also used for equivalence of |
|
2253 |
nonstandard ans standard definitions of pointwise |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2254 |
limit.*} |
|
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2255 |
|
| 14370 | 2256 |
(*----------------------------------------------------- |
2257 |
|X(n) - x| < 1/n ==> [<X n>] - hypreal_of_real x| \<in> Infinitesimal |
|
2258 |
-----------------------------------------------------*) |
|
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2259 |
lemma real_seq_to_hypreal_Infinitesimal: |
| 20563 | 2260 |
"\<forall>n. norm(X n - x) < inverse(real(Suc n)) |
2261 |
==> star_n X - star_of x \<in> Infinitesimal" |
|
2262 |
apply (auto intro!: bexI dest: FreeUltrafilterNat_inverse_real_of_posnat FreeUltrafilterNat_all FreeUltrafilterNat_Int intro: order_less_trans FreeUltrafilterNat_subset simp add: star_n_diff star_of_def Infinitesimal_FreeUltrafilterNat_iff star_n_inverse) |
|
| 14370 | 2263 |
done |
2264 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2265 |
lemma real_seq_to_hypreal_approx: |
| 20563 | 2266 |
"\<forall>n. norm(X n - x) < inverse(real(Suc n)) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2267 |
==> star_n X @= star_of x" |
| 14370 | 2268 |
apply (subst approx_minus_iff) |
2269 |
apply (rule mem_infmal_iff [THEN subst]) |
|
2270 |
apply (erule real_seq_to_hypreal_Infinitesimal) |
|
2271 |
done |
|
2272 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2273 |
lemma real_seq_to_hypreal_approx2: |
| 20563 | 2274 |
"\<forall>n. norm(x - X n) < inverse(real(Suc n)) |
|
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2275 |
==> star_n X @= star_of x" |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2276 |
apply (rule real_seq_to_hypreal_approx) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2277 |
apply (subst norm_minus_cancel [symmetric]) |
|
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20541
diff
changeset
|
2278 |
apply (simp del: norm_minus_cancel) |
| 14370 | 2279 |
done |
2280 |
||
|
14420
4e72cd222e0b
converted Hyperreal/HTranscendental to Isar script
paulson
parents:
14387
diff
changeset
|
2281 |
lemma real_seq_to_hypreal_Infinitesimal2: |
| 20563 | 2282 |
"\<forall>n. norm(X n - Y n) < inverse(real(Suc n)) |
2283 |
==> star_n X - star_n Y \<in> Infinitesimal" |
|
| 14370 | 2284 |
by (auto intro!: bexI |
2285 |
dest: FreeUltrafilterNat_inverse_real_of_posnat |
|
2286 |
FreeUltrafilterNat_all FreeUltrafilterNat_Int |
|
2287 |
intro: order_less_trans FreeUltrafilterNat_subset |
|
| 20563 | 2288 |
simp add: Infinitesimal_FreeUltrafilterNat_iff star_n_diff |
2289 |
star_n_inverse) |
|
| 14370 | 2290 |
|
| 10751 | 2291 |
end |