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