| author | wenzelm | 
| Fri, 14 Dec 2018 11:43:48 +0100 | |
| changeset 69470 | c8c3285f1294 | 
| parent 68527 | 2f4e2aab190a | 
| child 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 62479 | 1  | 
(* Title: HOL/Nonstandard_Analysis/NSA.thy  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
2  | 
Author: Jacques D. Fleuriot, University of Cambridge  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32155 
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, University of Cambridge  | 
| 27468 | 4  | 
*)  | 
5  | 
||
| 64435 | 6  | 
section \<open>Infinite Numbers, Infinitesimals, Infinitely Close Relation\<close>  | 
| 27468 | 7  | 
|
8  | 
theory NSA  | 
|
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
64438 
diff
changeset
 | 
9  | 
imports HyperDef "HOL-Library.Lub_Glb"  | 
| 27468 | 10  | 
begin  | 
11  | 
||
| 64435 | 12  | 
definition hnorm :: "'a::real_normed_vector star \<Rightarrow> real star"  | 
13  | 
where [transfer_unfold]: "hnorm = *f* norm"  | 
|
| 27468 | 14  | 
|
| 64435 | 15  | 
definition Infinitesimal  :: "('a::real_normed_vector) star set"
 | 
16  | 
  where "Infinitesimal = {x. \<forall>r \<in> Reals. 0 < r \<longrightarrow> hnorm x < r}"
 | 
|
| 27468 | 17  | 
|
| 64435 | 18  | 
definition HFinite :: "('a::real_normed_vector) star set"
 | 
19  | 
  where "HFinite = {x. \<exists>r \<in> Reals. hnorm x < r}"
 | 
|
| 27468 | 20  | 
|
| 64435 | 21  | 
definition HInfinite :: "('a::real_normed_vector) star set"
 | 
22  | 
  where "HInfinite = {x. \<forall>r \<in> Reals. r < hnorm x}"
 | 
|
| 27468 | 23  | 
|
| 64435 | 24  | 
definition approx :: "'a::real_normed_vector star \<Rightarrow> 'a star \<Rightarrow> bool" (infixl "\<approx>" 50)  | 
25  | 
where "x \<approx> y \<longleftrightarrow> x - y \<in> Infinitesimal"  | 
|
26  | 
\<comment> \<open>the ``infinitely close'' relation\<close>  | 
|
| 27468 | 27  | 
|
| 64435 | 28  | 
definition st :: "hypreal \<Rightarrow> hypreal"  | 
29  | 
where "st = (\<lambda>x. SOME r. x \<in> HFinite \<and> r \<in> \<real> \<and> r \<approx> x)"  | 
|
30  | 
\<comment> \<open>the standard part of a hyperreal\<close>  | 
|
| 27468 | 31  | 
|
| 64435 | 32  | 
definition monad :: "'a::real_normed_vector star \<Rightarrow> 'a star set"  | 
33  | 
  where "monad x = {y. x \<approx> y}"
 | 
|
| 27468 | 34  | 
|
| 64435 | 35  | 
definition galaxy :: "'a::real_normed_vector star \<Rightarrow> 'a star set"  | 
36  | 
  where "galaxy x = {y. (x + -y) \<in> HFinite}"
 | 
|
| 27468 | 37  | 
|
| 64435 | 38  | 
lemma SReal_def: "\<real> \<equiv> {x. \<exists>r. x = hypreal_of_real r}"
 | 
39  | 
by (simp add: Reals_def image_def)  | 
|
40  | 
||
| 27468 | 41  | 
|
| 61975 | 42  | 
subsection \<open>Nonstandard Extension of the Norm Function\<close>  | 
| 27468 | 43  | 
|
| 64435 | 44  | 
definition scaleHR :: "real star \<Rightarrow> 'a star \<Rightarrow> 'a::real_normed_vector star"  | 
45  | 
where [transfer_unfold]: "scaleHR = starfun2 scaleR"  | 
|
| 27468 | 46  | 
|
47  | 
lemma Standard_hnorm [simp]: "x \<in> Standard \<Longrightarrow> hnorm x \<in> Standard"  | 
|
| 64435 | 48  | 
by (simp add: hnorm_def)  | 
| 27468 | 49  | 
|
50  | 
lemma star_of_norm [simp]: "star_of (norm x) = hnorm (star_of x)"  | 
|
| 64435 | 51  | 
by transfer (rule refl)  | 
| 27468 | 52  | 
|
| 64435 | 53  | 
lemma hnorm_ge_zero [simp]: "\<And>x::'a::real_normed_vector star. 0 \<le> hnorm x"  | 
54  | 
by transfer (rule norm_ge_zero)  | 
|
| 27468 | 55  | 
|
| 64435 | 56  | 
lemma hnorm_eq_zero [simp]: "\<And>x::'a::real_normed_vector star. hnorm x = 0 \<longleftrightarrow> x = 0"  | 
57  | 
by transfer (rule norm_eq_zero)  | 
|
| 27468 | 58  | 
|
| 64435 | 59  | 
lemma hnorm_triangle_ineq: "\<And>x y::'a::real_normed_vector star. hnorm (x + y) \<le> hnorm x + hnorm y"  | 
60  | 
by transfer (rule norm_triangle_ineq)  | 
|
| 27468 | 61  | 
|
| 64435 | 62  | 
lemma hnorm_triangle_ineq3: "\<And>x y::'a::real_normed_vector star. \<bar>hnorm x - hnorm y\<bar> \<le> hnorm (x - y)"  | 
63  | 
by transfer (rule norm_triangle_ineq3)  | 
|
64  | 
||
65  | 
lemma hnorm_scaleR: "\<And>x::'a::real_normed_vector star. hnorm (a *\<^sub>R x) = \<bar>star_of a\<bar> * hnorm x"  | 
|
66  | 
by transfer (rule norm_scaleR)  | 
|
| 27468 | 67  | 
|
| 64435 | 68  | 
lemma hnorm_scaleHR: "\<And>a (x::'a::real_normed_vector star). hnorm (scaleHR a x) = \<bar>a\<bar> * hnorm x"  | 
69  | 
by transfer (rule norm_scaleR)  | 
|
| 27468 | 70  | 
|
| 64435 | 71  | 
lemma hnorm_mult_ineq: "\<And>x y::'a::real_normed_algebra star. hnorm (x * y) \<le> hnorm x * hnorm y"  | 
72  | 
by transfer (rule norm_mult_ineq)  | 
|
| 27468 | 73  | 
|
| 64435 | 74  | 
lemma hnorm_mult: "\<And>x y::'a::real_normed_div_algebra star. hnorm (x * y) = hnorm x * hnorm y"  | 
75  | 
by transfer (rule norm_mult)  | 
|
| 27468 | 76  | 
|
| 64435 | 77  | 
lemma hnorm_hyperpow: "\<And>(x::'a::{real_normed_div_algebra} star) n. hnorm (x pow n) = hnorm x pow n"
 | 
78  | 
by transfer (rule norm_power)  | 
|
79  | 
||
80  | 
lemma hnorm_one [simp]: "hnorm (1::'a::real_normed_div_algebra star) = 1"  | 
|
81  | 
by transfer (rule norm_one)  | 
|
| 27468 | 82  | 
|
| 64435 | 83  | 
lemma hnorm_zero [simp]: "hnorm (0::'a::real_normed_vector star) = 0"  | 
84  | 
by transfer (rule norm_zero)  | 
|
| 27468 | 85  | 
|
| 64435 | 86  | 
lemma zero_less_hnorm_iff [simp]: "\<And>x::'a::real_normed_vector star. 0 < hnorm x \<longleftrightarrow> x \<noteq> 0"  | 
87  | 
by transfer (rule zero_less_norm_iff)  | 
|
| 27468 | 88  | 
|
| 64435 | 89  | 
lemma hnorm_minus_cancel [simp]: "\<And>x::'a::real_normed_vector star. hnorm (- x) = hnorm x"  | 
90  | 
by transfer (rule norm_minus_cancel)  | 
|
| 27468 | 91  | 
|
| 64435 | 92  | 
lemma hnorm_minus_commute: "\<And>a b::'a::real_normed_vector star. hnorm (a - b) = hnorm (b - a)"  | 
93  | 
by transfer (rule norm_minus_commute)  | 
|
| 27468 | 94  | 
|
| 64435 | 95  | 
lemma hnorm_triangle_ineq2: "\<And>a b::'a::real_normed_vector star. hnorm a - hnorm b \<le> hnorm (a - b)"  | 
96  | 
by transfer (rule norm_triangle_ineq2)  | 
|
| 27468 | 97  | 
|
| 64435 | 98  | 
lemma hnorm_triangle_ineq4: "\<And>a b::'a::real_normed_vector star. hnorm (a - b) \<le> hnorm a + hnorm b"  | 
99  | 
by transfer (rule norm_triangle_ineq4)  | 
|
| 27468 | 100  | 
|
| 64435 | 101  | 
lemma abs_hnorm_cancel [simp]: "\<And>a::'a::real_normed_vector star. \<bar>hnorm a\<bar> = hnorm a"  | 
102  | 
by transfer (rule abs_norm_cancel)  | 
|
| 27468 | 103  | 
|
| 64435 | 104  | 
lemma hnorm_of_hypreal [simp]: "\<And>r. hnorm (of_hypreal r::'a::real_normed_algebra_1 star) = \<bar>r\<bar>"  | 
105  | 
by transfer (rule norm_of_real)  | 
|
| 27468 | 106  | 
|
107  | 
lemma nonzero_hnorm_inverse:  | 
|
| 64435 | 108  | 
"\<And>a::'a::real_normed_div_algebra star. a \<noteq> 0 \<Longrightarrow> hnorm (inverse a) = inverse (hnorm a)"  | 
109  | 
by transfer (rule nonzero_norm_inverse)  | 
|
| 27468 | 110  | 
|
111  | 
lemma hnorm_inverse:  | 
|
| 64435 | 112  | 
  "\<And>a::'a::{real_normed_div_algebra, division_ring} star. hnorm (inverse a) = inverse (hnorm a)"
 | 
113  | 
by transfer (rule norm_inverse)  | 
|
| 27468 | 114  | 
|
| 64435 | 115  | 
lemma hnorm_divide: "\<And>a b::'a::{real_normed_field, field} star. hnorm (a / b) = hnorm a / hnorm b"
 | 
116  | 
by transfer (rule norm_divide)  | 
|
| 27468 | 117  | 
|
| 64435 | 118  | 
lemma hypreal_hnorm_def [simp]: "\<And>r::hypreal. hnorm r = \<bar>r\<bar>"  | 
119  | 
by transfer (rule real_norm_def)  | 
|
| 27468 | 120  | 
|
121  | 
lemma hnorm_add_less:  | 
|
| 64435 | 122  | 
"\<And>(x::'a::real_normed_vector star) y r s. hnorm x < r \<Longrightarrow> hnorm y < s \<Longrightarrow> hnorm (x + y) < r + s"  | 
123  | 
by transfer (rule norm_add_less)  | 
|
| 27468 | 124  | 
|
125  | 
lemma hnorm_mult_less:  | 
|
| 64435 | 126  | 
"\<And>(x::'a::real_normed_algebra star) y r s. hnorm x < r \<Longrightarrow> hnorm y < s \<Longrightarrow> hnorm (x * y) < r * s"  | 
127  | 
by transfer (rule norm_mult_less)  | 
|
| 27468 | 128  | 
|
| 64435 | 129  | 
lemma hnorm_scaleHR_less: "\<bar>x\<bar> < r \<Longrightarrow> hnorm y < s \<Longrightarrow> hnorm (scaleHR x y) < r * s"  | 
130  | 
by (simp only: hnorm_scaleHR) (simp add: mult_strict_mono')  | 
|
131  | 
||
132  | 
||
133  | 
subsection \<open>Closure Laws for the Standard Reals\<close>  | 
|
| 27468 | 134  | 
|
| 64435 | 135  | 
lemma Reals_add_cancel: "x + y \<in> \<real> \<Longrightarrow> y \<in> \<real> \<Longrightarrow> x \<in> \<real>"  | 
136  | 
by (drule (1) Reals_diff) simp  | 
|
| 27468 | 137  | 
|
| 64435 | 138  | 
lemma SReal_hrabs: "x \<in> \<real> \<Longrightarrow> \<bar>x\<bar> \<in> \<real>"  | 
139  | 
for x :: hypreal  | 
|
140  | 
by (simp add: Reals_eq_Standard)  | 
|
| 27468 | 141  | 
|
| 61070 | 142  | 
lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \<in> \<real>"  | 
| 64435 | 143  | 
by (simp add: Reals_eq_Standard)  | 
| 27468 | 144  | 
|
| 64435 | 145  | 
lemma SReal_divide_numeral: "r \<in> \<real> \<Longrightarrow> r / (numeral w::hypreal) \<in> \<real>"  | 
146  | 
by simp  | 
|
| 27468 | 147  | 
|
| 61981 | 148  | 
text \<open>\<open>\<epsilon>\<close> is not in Reals because it is an infinitesimal\<close>  | 
149  | 
lemma SReal_epsilon_not_mem: "\<epsilon> \<notin> \<real>"  | 
|
| 64435 | 150  | 
by (auto simp: SReal_def hypreal_of_real_not_eq_epsilon [symmetric])  | 
| 27468 | 151  | 
|
| 61981 | 152  | 
lemma SReal_omega_not_mem: "\<omega> \<notin> \<real>"  | 
| 64435 | 153  | 
by (auto simp: SReal_def hypreal_of_real_not_eq_omega [symmetric])  | 
| 27468 | 154  | 
|
| 61070 | 155  | 
lemma SReal_UNIV_real: "{x. hypreal_of_real x \<in> \<real>} = (UNIV::real set)"
 | 
| 64435 | 156  | 
by simp  | 
| 27468 | 157  | 
|
| 64435 | 158  | 
lemma SReal_iff: "x \<in> \<real> \<longleftrightarrow> (\<exists>y. x = hypreal_of_real y)"  | 
159  | 
by (simp add: SReal_def)  | 
|
| 27468 | 160  | 
|
| 61070 | 161  | 
lemma hypreal_of_real_image: "hypreal_of_real `(UNIV::real set) = \<real>"  | 
| 64435 | 162  | 
by (simp add: Reals_eq_Standard Standard_def)  | 
| 27468 | 163  | 
|
| 61070 | 164  | 
lemma inv_hypreal_of_real_image: "inv hypreal_of_real ` \<real> = UNIV"  | 
| 64435 | 165  | 
apply (auto simp add: SReal_def)  | 
166  | 
apply (rule inj_star_of [THEN inv_f_f, THEN subst], blast)  | 
|
167  | 
done  | 
|
| 27468 | 168  | 
|
| 64435 | 169  | 
lemma SReal_hypreal_of_real_image: "\<exists>x. x \<in> P \<Longrightarrow> P \<subseteq> \<real> \<Longrightarrow> \<exists>Q. P = hypreal_of_real ` Q"  | 
170  | 
unfolding SReal_def image_def by blast  | 
|
| 27468 | 171  | 
|
| 64435 | 172  | 
lemma SReal_dense: "x \<in> \<real> \<Longrightarrow> y \<in> \<real> \<Longrightarrow> x < y \<Longrightarrow> \<exists>r \<in> Reals. x < r \<and> r < y"  | 
173  | 
for x y :: hypreal  | 
|
174  | 
apply (auto simp: SReal_def)  | 
|
175  | 
apply (drule dense)  | 
|
176  | 
apply auto  | 
|
177  | 
done  | 
|
| 27468 | 178  | 
|
| 64435 | 179  | 
|
180  | 
text \<open>Completeness of Reals, but both lemmas are unused.\<close>  | 
|
| 27468 | 181  | 
|
182  | 
lemma SReal_sup_lemma:  | 
|
| 64435 | 183  | 
"P \<subseteq> \<real> \<Longrightarrow> (\<exists>x \<in> P. y < x) = (\<exists>X. hypreal_of_real X \<in> P \<and> y < hypreal_of_real X)"  | 
184  | 
by (blast dest!: SReal_iff [THEN iffD1])  | 
|
| 27468 | 185  | 
|
186  | 
lemma SReal_sup_lemma2:  | 
|
| 64435 | 187  | 
"P \<subseteq> \<real> \<Longrightarrow> \<exists>x. x \<in> P \<Longrightarrow> \<exists>y \<in> Reals. \<forall>x \<in> P. x < y \<Longrightarrow>  | 
188  | 
    (\<exists>X. X \<in> {w. hypreal_of_real w \<in> P}) \<and>
 | 
|
189  | 
    (\<exists>Y. \<forall>X \<in> {w. hypreal_of_real w \<in> P}. X < Y)"
 | 
|
190  | 
apply (rule conjI)  | 
|
191  | 
apply (fast dest!: SReal_iff [THEN iffD1])  | 
|
192  | 
apply (auto, frule subsetD, assumption)  | 
|
193  | 
apply (drule SReal_iff [THEN iffD1])  | 
|
194  | 
apply (auto, rule_tac x = ya in exI, auto)  | 
|
195  | 
done  | 
|
| 27468 | 196  | 
|
197  | 
||
| 64435 | 198  | 
subsection \<open>Set of Finite Elements is a Subring of the Extended Reals\<close>  | 
| 27468 | 199  | 
|
| 64435 | 200  | 
lemma HFinite_add: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x + y \<in> HFinite"  | 
201  | 
unfolding HFinite_def by (blast intro!: Reals_add hnorm_add_less)  | 
|
| 27468 | 202  | 
|
| 64435 | 203  | 
lemma HFinite_mult: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x * y \<in> HFinite"  | 
204  | 
for x y :: "'a::real_normed_algebra star"  | 
|
205  | 
unfolding HFinite_def by (blast intro!: Reals_mult hnorm_mult_less)  | 
|
| 27468 | 206  | 
|
| 64435 | 207  | 
lemma HFinite_scaleHR: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> scaleHR x y \<in> HFinite"  | 
208  | 
by (auto simp: HFinite_def intro!: Reals_mult hnorm_scaleHR_less)  | 
|
| 27468 | 209  | 
|
| 64435 | 210  | 
lemma HFinite_minus_iff: "- x \<in> HFinite \<longleftrightarrow> x \<in> HFinite"  | 
211  | 
by (simp add: HFinite_def)  | 
|
| 27468 | 212  | 
|
213  | 
lemma HFinite_star_of [simp]: "star_of x \<in> HFinite"  | 
|
| 64435 | 214  | 
apply (simp add: HFinite_def)  | 
215  | 
apply (rule_tac x="star_of (norm x) + 1" in bexI)  | 
|
216  | 
apply (transfer, simp)  | 
|
217  | 
apply (blast intro: Reals_add SReal_hypreal_of_real Reals_1)  | 
|
218  | 
done  | 
|
| 27468 | 219  | 
|
| 61070 | 220  | 
lemma SReal_subset_HFinite: "(\<real>::hypreal set) \<subseteq> HFinite"  | 
| 64435 | 221  | 
by (auto simp add: SReal_def)  | 
| 27468 | 222  | 
|
| 64435 | 223  | 
lemma HFiniteD: "x \<in> HFinite \<Longrightarrow> \<exists>t \<in> Reals. hnorm x < t"  | 
224  | 
by (simp add: HFinite_def)  | 
|
| 27468 | 225  | 
|
| 64435 | 226  | 
lemma HFinite_hrabs_iff [iff]: "\<bar>x\<bar> \<in> HFinite \<longleftrightarrow> x \<in> HFinite"  | 
227  | 
for x :: hypreal  | 
|
228  | 
by (simp add: HFinite_def)  | 
|
| 27468 | 229  | 
|
| 64435 | 230  | 
lemma HFinite_hnorm_iff [iff]: "hnorm x \<in> HFinite \<longleftrightarrow> x \<in> HFinite"  | 
231  | 
for x :: hypreal  | 
|
232  | 
by (simp add: HFinite_def)  | 
|
| 27468 | 233  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
45541 
diff
changeset
 | 
234  | 
lemma HFinite_numeral [simp]: "numeral w \<in> HFinite"  | 
| 64435 | 235  | 
unfolding star_numeral_def by (rule HFinite_star_of)  | 
| 27468 | 236  | 
|
| 64435 | 237  | 
text \<open>As always with numerals, \<open>0\<close> and \<open>1\<close> are special cases.\<close>  | 
| 27468 | 238  | 
|
239  | 
lemma HFinite_0 [simp]: "0 \<in> HFinite"  | 
|
| 64435 | 240  | 
unfolding star_zero_def by (rule HFinite_star_of)  | 
| 27468 | 241  | 
|
242  | 
lemma HFinite_1 [simp]: "1 \<in> HFinite"  | 
|
| 64435 | 243  | 
unfolding star_one_def by (rule HFinite_star_of)  | 
| 27468 | 244  | 
|
| 64435 | 245  | 
lemma hrealpow_HFinite: "x \<in> HFinite \<Longrightarrow> x ^ n \<in> HFinite"  | 
246  | 
  for x :: "'a::{real_normed_algebra,monoid_mult} star"
 | 
|
247  | 
by (induct n) (auto simp add: power_Suc intro: HFinite_mult)  | 
|
| 27468 | 248  | 
|
| 64435 | 249  | 
lemma HFinite_bounded: "x \<in> HFinite \<Longrightarrow> y \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<in> HFinite"  | 
250  | 
for x y :: hypreal  | 
|
251  | 
apply (cases "x \<le> 0")  | 
|
252  | 
apply (drule_tac y = x in order_trans)  | 
|
253  | 
apply (drule_tac [2] order_antisym)  | 
|
254  | 
apply (auto simp add: linorder_not_le)  | 
|
255  | 
apply (auto intro: order_le_less_trans simp add: abs_if HFinite_def)  | 
|
256  | 
done  | 
|
| 27468 | 257  | 
|
258  | 
||
| 64435 | 259  | 
subsection \<open>Set of Infinitesimals is a Subring of the Hyperreals\<close>  | 
| 27468 | 260  | 
|
| 64435 | 261  | 
lemma InfinitesimalI: "(\<And>r. r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> hnorm x < r) \<Longrightarrow> x \<in> Infinitesimal"  | 
262  | 
by (simp add: Infinitesimal_def)  | 
|
| 27468 | 263  | 
|
| 64435 | 264  | 
lemma InfinitesimalD: "x \<in> Infinitesimal \<Longrightarrow> \<forall>r \<in> Reals. 0 < r \<longrightarrow> hnorm x < r"  | 
265  | 
by (simp add: Infinitesimal_def)  | 
|
| 27468 | 266  | 
|
| 64435 | 267  | 
lemma InfinitesimalI2: "(\<And>r. 0 < r \<Longrightarrow> hnorm x < star_of r) \<Longrightarrow> x \<in> Infinitesimal"  | 
268  | 
by (auto simp add: Infinitesimal_def SReal_def)  | 
|
| 27468 | 269  | 
|
| 64435 | 270  | 
lemma InfinitesimalD2: "x \<in> Infinitesimal \<Longrightarrow> 0 < r \<Longrightarrow> hnorm x < star_of r"  | 
271  | 
by (auto simp add: Infinitesimal_def SReal_def)  | 
|
| 27468 | 272  | 
|
273  | 
lemma Infinitesimal_zero [iff]: "0 \<in> Infinitesimal"  | 
|
| 64435 | 274  | 
by (simp add: Infinitesimal_def)  | 
| 27468 | 275  | 
|
| 64435 | 276  | 
lemma Infinitesimal_add: "x \<in> Infinitesimal \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x + y \<in> Infinitesimal"  | 
277  | 
apply (rule InfinitesimalI)  | 
|
| 
68527
 
2f4e2aab190a
Generalising and renaming some basic results
 
paulson <lp15@cam.ac.uk> 
parents: 
68501 
diff
changeset
 | 
278  | 
apply (rule field_sum_of_halves [THEN subst])  | 
| 64435 | 279  | 
apply (drule half_gt_zero)  | 
280  | 
apply (blast intro: hnorm_add_less SReal_divide_numeral dest: InfinitesimalD)  | 
|
281  | 
done  | 
|
| 27468 | 282  | 
|
| 64435 | 283  | 
lemma Infinitesimal_minus_iff [simp]: "- x \<in> Infinitesimal \<longleftrightarrow> x \<in> Infinitesimal"  | 
284  | 
by (simp add: Infinitesimal_def)  | 
|
| 27468 | 285  | 
|
| 64435 | 286  | 
lemma Infinitesimal_hnorm_iff: "hnorm x \<in> Infinitesimal \<longleftrightarrow> x \<in> Infinitesimal"  | 
287  | 
by (simp add: Infinitesimal_def)  | 
|
| 27468 | 288  | 
|
| 64435 | 289  | 
lemma Infinitesimal_hrabs_iff [iff]: "\<bar>x\<bar> \<in> Infinitesimal \<longleftrightarrow> x \<in> Infinitesimal"  | 
290  | 
for x :: hypreal  | 
|
291  | 
by (simp add: abs_if)  | 
|
| 27468 | 292  | 
|
293  | 
lemma Infinitesimal_of_hypreal_iff [simp]:  | 
|
| 64435 | 294  | 
"(of_hypreal x::'a::real_normed_algebra_1 star) \<in> Infinitesimal \<longleftrightarrow> x \<in> Infinitesimal"  | 
295  | 
by (subst Infinitesimal_hnorm_iff [symmetric]) simp  | 
|
| 27468 | 296  | 
|
| 64435 | 297  | 
lemma Infinitesimal_diff: "x \<in> Infinitesimal \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x - y \<in> Infinitesimal"  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
51521 
diff
changeset
 | 
298  | 
using Infinitesimal_add [of x "- y"] by simp  | 
| 27468 | 299  | 
|
| 64435 | 300  | 
lemma Infinitesimal_mult: "x \<in> Infinitesimal \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x * y \<in> Infinitesimal"  | 
301  | 
for x y :: "'a::real_normed_algebra star"  | 
|
302  | 
apply (rule InfinitesimalI)  | 
|
303  | 
apply (subgoal_tac "hnorm (x * y) < 1 * r")  | 
|
304  | 
apply (simp only: mult_1)  | 
|
305  | 
apply (rule hnorm_mult_less)  | 
|
306  | 
apply (simp_all add: InfinitesimalD)  | 
|
307  | 
done  | 
|
| 27468 | 308  | 
|
| 64435 | 309  | 
lemma Infinitesimal_HFinite_mult: "x \<in> Infinitesimal \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x * y \<in> Infinitesimal"  | 
310  | 
for x y :: "'a::real_normed_algebra star"  | 
|
311  | 
apply (rule InfinitesimalI)  | 
|
312  | 
apply (drule HFiniteD, clarify)  | 
|
313  | 
apply (subgoal_tac "0 < t")  | 
|
314  | 
apply (subgoal_tac "hnorm (x * y) < (r / t) * t", simp)  | 
|
315  | 
apply (subgoal_tac "0 < r / t")  | 
|
316  | 
apply (rule hnorm_mult_less)  | 
|
317  | 
apply (simp add: InfinitesimalD)  | 
|
318  | 
apply assumption  | 
|
319  | 
apply simp  | 
|
320  | 
apply (erule order_le_less_trans [OF hnorm_ge_zero])  | 
|
321  | 
done  | 
|
| 27468 | 322  | 
|
323  | 
lemma Infinitesimal_HFinite_scaleHR:  | 
|
| 64435 | 324  | 
"x \<in> Infinitesimal \<Longrightarrow> y \<in> HFinite \<Longrightarrow> scaleHR x y \<in> Infinitesimal"  | 
325  | 
apply (rule InfinitesimalI)  | 
|
326  | 
apply (drule HFiniteD, clarify)  | 
|
327  | 
apply (drule InfinitesimalD)  | 
|
328  | 
apply (simp add: hnorm_scaleHR)  | 
|
329  | 
apply (subgoal_tac "0 < t")  | 
|
330  | 
apply (subgoal_tac "\<bar>x\<bar> * hnorm y < (r / t) * t", simp)  | 
|
331  | 
apply (subgoal_tac "0 < r / t")  | 
|
332  | 
apply (rule mult_strict_mono', simp_all)  | 
|
333  | 
apply (erule order_le_less_trans [OF hnorm_ge_zero])  | 
|
334  | 
done  | 
|
| 27468 | 335  | 
|
336  | 
lemma Infinitesimal_HFinite_mult2:  | 
|
| 64435 | 337  | 
"x \<in> Infinitesimal \<Longrightarrow> y \<in> HFinite \<Longrightarrow> y * x \<in> Infinitesimal"  | 
338  | 
for x y :: "'a::real_normed_algebra star"  | 
|
339  | 
apply (rule InfinitesimalI)  | 
|
340  | 
apply (drule HFiniteD, clarify)  | 
|
341  | 
apply (subgoal_tac "0 < t")  | 
|
342  | 
apply (subgoal_tac "hnorm (y * x) < t * (r / t)", simp)  | 
|
343  | 
apply (subgoal_tac "0 < r / t")  | 
|
344  | 
apply (rule hnorm_mult_less)  | 
|
345  | 
apply assumption  | 
|
346  | 
apply (simp add: InfinitesimalD)  | 
|
347  | 
apply simp  | 
|
348  | 
apply (erule order_le_less_trans [OF hnorm_ge_zero])  | 
|
349  | 
done  | 
|
| 27468 | 350  | 
|
| 64435 | 351  | 
lemma Infinitesimal_scaleR2: "x \<in> Infinitesimal \<Longrightarrow> a *\<^sub>R x \<in> Infinitesimal"  | 
352  | 
apply (case_tac "a = 0", simp)  | 
|
353  | 
apply (rule InfinitesimalI)  | 
|
354  | 
apply (drule InfinitesimalD)  | 
|
355  | 
apply (drule_tac x="r / \<bar>star_of a\<bar>" in bspec)  | 
|
356  | 
apply (simp add: Reals_eq_Standard)  | 
|
357  | 
apply simp  | 
|
358  | 
apply (simp add: hnorm_scaleR pos_less_divide_eq mult.commute)  | 
|
359  | 
done  | 
|
| 27468 | 360  | 
|
361  | 
lemma Compl_HFinite: "- HFinite = HInfinite"  | 
|
| 64435 | 362  | 
apply (auto simp add: HInfinite_def HFinite_def linorder_not_less)  | 
363  | 
apply (rule_tac y="r + 1" in order_less_le_trans, simp)  | 
|
364  | 
apply simp  | 
|
365  | 
done  | 
|
| 27468 | 366  | 
|
| 64435 | 367  | 
lemma HInfinite_inverse_Infinitesimal: "x \<in> HInfinite \<Longrightarrow> inverse x \<in> Infinitesimal"  | 
368  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
369  | 
apply (rule InfinitesimalI)  | 
|
370  | 
apply (subgoal_tac "x \<noteq> 0")  | 
|
371  | 
apply (rule inverse_less_imp_less)  | 
|
372  | 
apply (simp add: nonzero_hnorm_inverse)  | 
|
373  | 
apply (simp add: HInfinite_def Reals_inverse)  | 
|
374  | 
apply assumption  | 
|
375  | 
apply (clarify, simp add: Compl_HFinite [symmetric])  | 
|
376  | 
done  | 
|
| 27468 | 377  | 
|
378  | 
lemma HInfiniteI: "(\<And>r. r \<in> \<real> \<Longrightarrow> r < hnorm x) \<Longrightarrow> x \<in> HInfinite"  | 
|
| 64435 | 379  | 
by (simp add: HInfinite_def)  | 
| 27468 | 380  | 
|
| 64435 | 381  | 
lemma HInfiniteD: "x \<in> HInfinite \<Longrightarrow> r \<in> \<real> \<Longrightarrow> r < hnorm x"  | 
382  | 
by (simp add: HInfinite_def)  | 
|
| 27468 | 383  | 
|
| 64435 | 384  | 
lemma HInfinite_mult: "x \<in> HInfinite \<Longrightarrow> y \<in> HInfinite \<Longrightarrow> x * y \<in> HInfinite"  | 
385  | 
for x y :: "'a::real_normed_div_algebra star"  | 
|
386  | 
apply (rule HInfiniteI, simp only: hnorm_mult)  | 
|
387  | 
apply (subgoal_tac "r * 1 < hnorm x * hnorm y", simp only: mult_1)  | 
|
388  | 
apply (case_tac "x = 0", simp add: HInfinite_def)  | 
|
389  | 
apply (rule mult_strict_mono)  | 
|
390  | 
apply (simp_all add: HInfiniteD)  | 
|
391  | 
done  | 
|
| 27468 | 392  | 
|
| 64435 | 393  | 
lemma hypreal_add_zero_less_le_mono: "r < x \<Longrightarrow> 0 \<le> y \<Longrightarrow> r < x + y"  | 
394  | 
for r x y :: hypreal  | 
|
395  | 
by (auto dest: add_less_le_mono)  | 
|
| 27468 | 396  | 
|
| 64435 | 397  | 
lemma HInfinite_add_ge_zero: "x \<in> HInfinite \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> x + y \<in> HInfinite"  | 
398  | 
for x y :: hypreal  | 
|
399  | 
by (auto simp: abs_if add.commute HInfinite_def)  | 
|
| 27468 | 400  | 
|
| 64435 | 401  | 
lemma HInfinite_add_ge_zero2: "x \<in> HInfinite \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> y + x \<in> HInfinite"  | 
402  | 
for x y :: hypreal  | 
|
403  | 
by (auto intro!: HInfinite_add_ge_zero simp add: add.commute)  | 
|
| 27468 | 404  | 
|
| 64435 | 405  | 
lemma HInfinite_add_gt_zero: "x \<in> HInfinite \<Longrightarrow> 0 < y \<Longrightarrow> 0 < x \<Longrightarrow> x + y \<in> HInfinite"  | 
406  | 
for x y :: hypreal  | 
|
407  | 
by (blast intro: HInfinite_add_ge_zero order_less_imp_le)  | 
|
| 27468 | 408  | 
|
| 64435 | 409  | 
lemma HInfinite_minus_iff: "- x \<in> HInfinite \<longleftrightarrow> x \<in> HInfinite"  | 
410  | 
by (simp add: HInfinite_def)  | 
|
| 27468 | 411  | 
|
| 64435 | 412  | 
lemma HInfinite_add_le_zero: "x \<in> HInfinite \<Longrightarrow> y \<le> 0 \<Longrightarrow> x \<le> 0 \<Longrightarrow> x + y \<in> HInfinite"  | 
413  | 
for x y :: hypreal  | 
|
414  | 
apply (drule HInfinite_minus_iff [THEN iffD2])  | 
|
415  | 
apply (rule HInfinite_minus_iff [THEN iffD1])  | 
|
416  | 
apply (simp only: minus_add add.commute)  | 
|
417  | 
apply (rule HInfinite_add_ge_zero)  | 
|
418  | 
apply simp_all  | 
|
419  | 
done  | 
|
| 27468 | 420  | 
|
| 64435 | 421  | 
lemma HInfinite_add_lt_zero: "x \<in> HInfinite \<Longrightarrow> y < 0 \<Longrightarrow> x < 0 \<Longrightarrow> x + y \<in> HInfinite"  | 
422  | 
for x y :: hypreal  | 
|
423  | 
by (blast intro: HInfinite_add_le_zero order_less_imp_le)  | 
|
| 27468 | 424  | 
|
425  | 
lemma HFinite_sum_squares:  | 
|
| 64435 | 426  | 
"a \<in> HFinite \<Longrightarrow> b \<in> HFinite \<Longrightarrow> c \<in> HFinite \<Longrightarrow> a * a + b * b + c * c \<in> HFinite"  | 
427  | 
for a b c :: "'a::real_normed_algebra star"  | 
|
428  | 
by (auto intro: HFinite_mult HFinite_add)  | 
|
| 27468 | 429  | 
|
| 64435 | 430  | 
lemma not_Infinitesimal_not_zero: "x \<notin> Infinitesimal \<Longrightarrow> x \<noteq> 0"  | 
431  | 
by auto  | 
|
| 27468 | 432  | 
|
| 64435 | 433  | 
lemma not_Infinitesimal_not_zero2: "x \<in> HFinite - Infinitesimal \<Longrightarrow> x \<noteq> 0"  | 
434  | 
by auto  | 
|
| 27468 | 435  | 
|
436  | 
lemma HFinite_diff_Infinitesimal_hrabs:  | 
|
| 64435 | 437  | 
"x \<in> HFinite - Infinitesimal \<Longrightarrow> \<bar>x\<bar> \<in> HFinite - Infinitesimal"  | 
438  | 
for x :: hypreal  | 
|
439  | 
by blast  | 
|
| 27468 | 440  | 
|
| 64435 | 441  | 
lemma hnorm_le_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> hnorm x \<le> e \<Longrightarrow> x \<in> Infinitesimal"  | 
442  | 
by (auto simp: Infinitesimal_def abs_less_iff)  | 
|
| 27468 | 443  | 
|
| 64435 | 444  | 
lemma hnorm_less_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> hnorm x < e \<Longrightarrow> x \<in> Infinitesimal"  | 
445  | 
by (erule hnorm_le_Infinitesimal, erule order_less_imp_le)  | 
|
| 27468 | 446  | 
|
| 64435 | 447  | 
lemma hrabs_le_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> \<bar>x\<bar> \<le> e \<Longrightarrow> x \<in> Infinitesimal"  | 
448  | 
for x :: hypreal  | 
|
449  | 
by (erule hnorm_le_Infinitesimal) simp  | 
|
| 27468 | 450  | 
|
| 64435 | 451  | 
lemma hrabs_less_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> \<bar>x\<bar> < e \<Longrightarrow> x \<in> Infinitesimal"  | 
452  | 
for x :: hypreal  | 
|
453  | 
by (erule hnorm_less_Infinitesimal) simp  | 
|
| 27468 | 454  | 
|
455  | 
lemma Infinitesimal_interval:  | 
|
| 64435 | 456  | 
"e \<in> Infinitesimal \<Longrightarrow> e' \<in> Infinitesimal \<Longrightarrow> e' < x \<Longrightarrow> x < e \<Longrightarrow> x \<in> Infinitesimal"  | 
457  | 
for x :: hypreal  | 
|
458  | 
by (auto simp add: Infinitesimal_def abs_less_iff)  | 
|
| 27468 | 459  | 
|
460  | 
lemma Infinitesimal_interval2:  | 
|
| 64435 | 461  | 
"e \<in> Infinitesimal \<Longrightarrow> e' \<in> Infinitesimal \<Longrightarrow> e' \<le> x \<Longrightarrow> x \<le> e \<Longrightarrow> x \<in> Infinitesimal"  | 
462  | 
for x :: hypreal  | 
|
463  | 
by (auto intro: Infinitesimal_interval simp add: order_le_less)  | 
|
| 27468 | 464  | 
|
465  | 
||
| 64435 | 466  | 
lemma lemma_Infinitesimal_hyperpow: "x \<in> Infinitesimal \<Longrightarrow> 0 < N \<Longrightarrow> \<bar>x pow N\<bar> \<le> \<bar>x\<bar>"  | 
467  | 
for x :: hypreal  | 
|
468  | 
apply (unfold Infinitesimal_def)  | 
|
469  | 
apply (auto intro!: hyperpow_Suc_le_self2  | 
|
470  | 
simp: hyperpow_hrabs [symmetric] hypnat_gt_zero_iff2 abs_ge_zero)  | 
|
471  | 
done  | 
|
| 27468 | 472  | 
|
| 64435 | 473  | 
lemma Infinitesimal_hyperpow: "x \<in> Infinitesimal \<Longrightarrow> 0 < N \<Longrightarrow> x pow N \<in> Infinitesimal"  | 
474  | 
for x :: hypreal  | 
|
475  | 
apply (rule hrabs_le_Infinitesimal)  | 
|
476  | 
apply (rule_tac [2] lemma_Infinitesimal_hyperpow)  | 
|
477  | 
apply auto  | 
|
478  | 
done  | 
|
| 27468 | 479  | 
|
480  | 
lemma hrealpow_hyperpow_Infinitesimal_iff:  | 
|
| 64435 | 481  | 
"(x ^ n \<in> Infinitesimal) \<longleftrightarrow> x pow (hypnat_of_nat n) \<in> Infinitesimal"  | 
482  | 
by (simp only: hyperpow_hypnat_of_nat)  | 
|
| 27468 | 483  | 
|
| 64435 | 484  | 
lemma Infinitesimal_hrealpow: "x \<in> Infinitesimal \<Longrightarrow> 0 < n \<Longrightarrow> x ^ n \<in> Infinitesimal"  | 
485  | 
for x :: hypreal  | 
|
486  | 
by (simp add: hrealpow_hyperpow_Infinitesimal_iff Infinitesimal_hyperpow)  | 
|
| 27468 | 487  | 
|
488  | 
lemma not_Infinitesimal_mult:  | 
|
| 64435 | 489  | 
"x \<notin> Infinitesimal \<Longrightarrow> y \<notin> Infinitesimal \<Longrightarrow> x * y \<notin> Infinitesimal"  | 
490  | 
for x y :: "'a::real_normed_div_algebra star"  | 
|
491  | 
apply (unfold Infinitesimal_def, clarify, rename_tac r s)  | 
|
492  | 
apply (simp only: linorder_not_less hnorm_mult)  | 
|
493  | 
apply (drule_tac x = "r * s" in bspec)  | 
|
494  | 
apply (fast intro: Reals_mult)  | 
|
495  | 
apply simp  | 
|
496  | 
apply (drule_tac c = s and d = "hnorm y" and a = r and b = "hnorm x" in mult_mono)  | 
|
497  | 
apply simp_all  | 
|
498  | 
done  | 
|
| 27468 | 499  | 
|
| 64435 | 500  | 
lemma Infinitesimal_mult_disj: "x * y \<in> Infinitesimal \<Longrightarrow> x \<in> Infinitesimal \<or> y \<in> Infinitesimal"  | 
501  | 
for x y :: "'a::real_normed_div_algebra star"  | 
|
502  | 
apply (rule ccontr)  | 
|
503  | 
apply (drule de_Morgan_disj [THEN iffD1])  | 
|
504  | 
apply (fast dest: not_Infinitesimal_mult)  | 
|
505  | 
done  | 
|
| 27468 | 506  | 
|
| 64435 | 507  | 
lemma HFinite_Infinitesimal_not_zero: "x \<in> HFinite-Infinitesimal \<Longrightarrow> x \<noteq> 0"  | 
508  | 
by blast  | 
|
| 27468 | 509  | 
|
510  | 
lemma HFinite_Infinitesimal_diff_mult:  | 
|
| 64435 | 511  | 
"x \<in> HFinite - Infinitesimal \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> x * y \<in> HFinite - Infinitesimal"  | 
512  | 
for x y :: "'a::real_normed_div_algebra star"  | 
|
513  | 
apply clarify  | 
|
514  | 
apply (blast dest: HFinite_mult not_Infinitesimal_mult)  | 
|
515  | 
done  | 
|
| 27468 | 516  | 
|
| 64435 | 517  | 
lemma Infinitesimal_subset_HFinite: "Infinitesimal \<subseteq> HFinite"  | 
518  | 
apply (simp add: Infinitesimal_def HFinite_def)  | 
|
519  | 
apply auto  | 
|
520  | 
apply (rule_tac x = 1 in bexI)  | 
|
521  | 
apply auto  | 
|
522  | 
done  | 
|
| 27468 | 523  | 
|
| 64435 | 524  | 
lemma Infinitesimal_star_of_mult: "x \<in> Infinitesimal \<Longrightarrow> x * star_of r \<in> Infinitesimal"  | 
525  | 
for x :: "'a::real_normed_algebra star"  | 
|
526  | 
by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult])  | 
|
| 27468 | 527  | 
|
| 64435 | 528  | 
lemma Infinitesimal_star_of_mult2: "x \<in> Infinitesimal \<Longrightarrow> star_of r * x \<in> Infinitesimal"  | 
529  | 
for x :: "'a::real_normed_algebra star"  | 
|
530  | 
by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult2])  | 
|
| 27468 | 531  | 
|
532  | 
||
| 64435 | 533  | 
subsection \<open>The Infinitely Close Relation\<close>  | 
| 27468 | 534  | 
|
| 64435 | 535  | 
lemma mem_infmal_iff: "x \<in> Infinitesimal \<longleftrightarrow> x \<approx> 0"  | 
536  | 
by (simp add: Infinitesimal_def approx_def)  | 
|
| 27468 | 537  | 
|
| 64435 | 538  | 
lemma approx_minus_iff: "x \<approx> y \<longleftrightarrow> x - y \<approx> 0"  | 
539  | 
by (simp add: approx_def)  | 
|
| 27468 | 540  | 
|
| 64435 | 541  | 
lemma approx_minus_iff2: "x \<approx> y \<longleftrightarrow> - y + x \<approx> 0"  | 
542  | 
by (simp add: approx_def add.commute)  | 
|
| 27468 | 543  | 
|
| 61982 | 544  | 
lemma approx_refl [iff]: "x \<approx> x"  | 
| 64435 | 545  | 
by (simp add: approx_def Infinitesimal_def)  | 
| 27468 | 546  | 
|
| 64435 | 547  | 
lemma hypreal_minus_distrib1: "- (y + - x) = x + -y"  | 
548  | 
for x y :: "'a::ab_group_add"  | 
|
549  | 
by (simp add: add.commute)  | 
|
| 27468 | 550  | 
|
| 64435 | 551  | 
lemma approx_sym: "x \<approx> y \<Longrightarrow> y \<approx> x"  | 
552  | 
apply (simp add: approx_def)  | 
|
553  | 
apply (drule Infinitesimal_minus_iff [THEN iffD2])  | 
|
554  | 
apply simp  | 
|
555  | 
done  | 
|
| 27468 | 556  | 
|
| 64435 | 557  | 
lemma approx_trans: "x \<approx> y \<Longrightarrow> y \<approx> z \<Longrightarrow> x \<approx> z"  | 
558  | 
apply (simp add: approx_def)  | 
|
559  | 
apply (drule (1) Infinitesimal_add)  | 
|
560  | 
apply simp  | 
|
561  | 
done  | 
|
| 27468 | 562  | 
|
| 64435 | 563  | 
lemma approx_trans2: "r \<approx> x \<Longrightarrow> s \<approx> x \<Longrightarrow> r \<approx> s"  | 
564  | 
by (blast intro: approx_sym approx_trans)  | 
|
| 27468 | 565  | 
|
| 64435 | 566  | 
lemma approx_trans3: "x \<approx> r \<Longrightarrow> x \<approx> s \<Longrightarrow> r \<approx> s"  | 
567  | 
by (blast intro: approx_sym approx_trans)  | 
|
| 27468 | 568  | 
|
| 64435 | 569  | 
lemma approx_reorient: "x \<approx> y \<longleftrightarrow> y \<approx> x"  | 
570  | 
by (blast intro: approx_sym)  | 
|
| 27468 | 571  | 
|
| 64435 | 572  | 
text \<open>Reorientation simplification procedure: reorients (polymorphic)  | 
573  | 
\<open>0 = x\<close>, \<open>1 = x\<close>, \<open>nnn = x\<close> provided \<open>x\<close> isn't \<open>0\<close>, \<open>1\<close> or a numeral.\<close>  | 
|
| 
45541
 
934866fc776c
simplify implementation of approx_reorient_simproc
 
huffman 
parents: 
45540 
diff
changeset
 | 
574  | 
simproc_setup approx_reorient_simproc  | 
| 61982 | 575  | 
  ("0 \<approx> x" | "1 \<approx> y" | "numeral w \<approx> z" | "- 1 \<approx> y" | "- numeral w \<approx> r") =
 | 
| 61975 | 576  | 
\<open>  | 
| 
45541
 
934866fc776c
simplify implementation of approx_reorient_simproc
 
huffman 
parents: 
45540 
diff
changeset
 | 
577  | 
  let val rule = @{thm approx_reorient} RS eq_reflection
 | 
| 59582 | 578  | 
fun proc phi ss ct =  | 
579  | 
case Thm.term_of ct of  | 
|
| 
45541
 
934866fc776c
simplify implementation of approx_reorient_simproc
 
huffman 
parents: 
45540 
diff
changeset
 | 
580  | 
_ $ t $ u => if can HOLogic.dest_number u then NONE  | 
| 
 
934866fc776c
simplify implementation of approx_reorient_simproc
 
huffman 
parents: 
45540 
diff
changeset
 | 
581  | 
else if can HOLogic.dest_number t then SOME rule else NONE  | 
| 
 
934866fc776c
simplify implementation of approx_reorient_simproc
 
huffman 
parents: 
45540 
diff
changeset
 | 
582  | 
| _ => NONE  | 
| 
 
934866fc776c
simplify implementation of approx_reorient_simproc
 
huffman 
parents: 
45540 
diff
changeset
 | 
583  | 
in proc end  | 
| 61975 | 584  | 
\<close>  | 
| 27468 | 585  | 
|
| 64435 | 586  | 
lemma Infinitesimal_approx_minus: "x - y \<in> Infinitesimal \<longleftrightarrow> x \<approx> y"  | 
587  | 
by (simp add: approx_minus_iff [symmetric] mem_infmal_iff)  | 
|
| 27468 | 588  | 
|
| 64435 | 589  | 
lemma approx_monad_iff: "x \<approx> y \<longleftrightarrow> monad x = monad y"  | 
590  | 
by (auto simp add: monad_def dest: approx_sym elim!: approx_trans equalityCE)  | 
|
| 27468 | 591  | 
|
| 64435 | 592  | 
lemma Infinitesimal_approx: "x \<in> Infinitesimal \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x \<approx> y"  | 
593  | 
apply (simp add: mem_infmal_iff)  | 
|
594  | 
apply (blast intro: approx_trans approx_sym)  | 
|
595  | 
done  | 
|
| 27468 | 596  | 
|
| 64435 | 597  | 
lemma approx_add: "a \<approx> b \<Longrightarrow> c \<approx> d \<Longrightarrow> a + c \<approx> b + d"  | 
| 27468 | 598  | 
proof (unfold approx_def)  | 
599  | 
assume inf: "a - b \<in> Infinitesimal" "c - d \<in> Infinitesimal"  | 
|
600  | 
have "a + c - (b + d) = (a - b) + (c - d)" by simp  | 
|
| 64435 | 601  | 
also have "... \<in> Infinitesimal"  | 
602  | 
using inf by (rule Infinitesimal_add)  | 
|
| 27468 | 603  | 
finally show "a + c - (b + d) \<in> Infinitesimal" .  | 
604  | 
qed  | 
|
605  | 
||
| 64435 | 606  | 
lemma approx_minus: "a \<approx> b \<Longrightarrow> - a \<approx> - b"  | 
607  | 
apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])  | 
|
608  | 
apply (drule approx_minus_iff [THEN iffD1])  | 
|
609  | 
apply (simp add: add.commute)  | 
|
610  | 
done  | 
|
| 27468 | 611  | 
|
| 64435 | 612  | 
lemma approx_minus2: "- a \<approx> - b \<Longrightarrow> a \<approx> b"  | 
613  | 
by (auto dest: approx_minus)  | 
|
| 27468 | 614  | 
|
| 64435 | 615  | 
lemma approx_minus_cancel [simp]: "- a \<approx> - b \<longleftrightarrow> a \<approx> b"  | 
616  | 
by (blast intro: approx_minus approx_minus2)  | 
|
| 27468 | 617  | 
|
| 64435 | 618  | 
lemma approx_add_minus: "a \<approx> b \<Longrightarrow> c \<approx> d \<Longrightarrow> a + - c \<approx> b + - d"  | 
619  | 
by (blast intro!: approx_add approx_minus)  | 
|
| 27468 | 620  | 
|
| 64435 | 621  | 
lemma approx_diff: "a \<approx> b \<Longrightarrow> c \<approx> d \<Longrightarrow> a - c \<approx> b - d"  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
51521 
diff
changeset
 | 
622  | 
using approx_add [of a b "- c" "- d"] by simp  | 
| 27468 | 623  | 
|
| 64435 | 624  | 
lemma approx_mult1: "a \<approx> b \<Longrightarrow> c \<in> HFinite \<Longrightarrow> a * c \<approx> b * c"  | 
625  | 
for a b c :: "'a::real_normed_algebra star"  | 
|
626  | 
by (simp add: approx_def Infinitesimal_HFinite_mult left_diff_distrib [symmetric])  | 
|
627  | 
||
628  | 
lemma approx_mult2: "a \<approx> b \<Longrightarrow> c \<in> HFinite \<Longrightarrow> c * a \<approx> c * b"  | 
|
629  | 
for a b c :: "'a::real_normed_algebra star"  | 
|
630  | 
by (simp add: approx_def Infinitesimal_HFinite_mult2 right_diff_distrib [symmetric])  | 
|
| 27468 | 631  | 
|
| 64435 | 632  | 
lemma approx_mult_subst: "u \<approx> v * x \<Longrightarrow> x \<approx> y \<Longrightarrow> v \<in> HFinite \<Longrightarrow> u \<approx> v * y"  | 
633  | 
for u v x y :: "'a::real_normed_algebra star"  | 
|
634  | 
by (blast intro: approx_mult2 approx_trans)  | 
|
| 27468 | 635  | 
|
| 64435 | 636  | 
lemma approx_mult_subst2: "u \<approx> x * v \<Longrightarrow> x \<approx> y \<Longrightarrow> v \<in> HFinite \<Longrightarrow> u \<approx> y * v"  | 
637  | 
for u v x y :: "'a::real_normed_algebra star"  | 
|
638  | 
by (blast intro: approx_mult1 approx_trans)  | 
|
| 27468 | 639  | 
|
| 64435 | 640  | 
lemma approx_mult_subst_star_of: "u \<approx> x * star_of v \<Longrightarrow> x \<approx> y \<Longrightarrow> u \<approx> y * star_of v"  | 
641  | 
for u x y :: "'a::real_normed_algebra star"  | 
|
642  | 
by (auto intro: approx_mult_subst2)  | 
|
| 27468 | 643  | 
|
| 64435 | 644  | 
lemma approx_eq_imp: "a = b \<Longrightarrow> a \<approx> b"  | 
645  | 
by (simp add: approx_def)  | 
|
| 27468 | 646  | 
|
| 64435 | 647  | 
lemma Infinitesimal_minus_approx: "x \<in> Infinitesimal \<Longrightarrow> - x \<approx> x"  | 
648  | 
by (blast intro: Infinitesimal_minus_iff [THEN iffD2] mem_infmal_iff [THEN iffD1] approx_trans2)  | 
|
| 27468 | 649  | 
|
| 64435 | 650  | 
lemma bex_Infinitesimal_iff: "(\<exists>y \<in> Infinitesimal. x - z = y) \<longleftrightarrow> x \<approx> z"  | 
651  | 
by (simp add: approx_def)  | 
|
| 27468 | 652  | 
|
| 64435 | 653  | 
lemma bex_Infinitesimal_iff2: "(\<exists>y \<in> Infinitesimal. x = z + y) \<longleftrightarrow> x \<approx> z"  | 
654  | 
by (force simp add: bex_Infinitesimal_iff [symmetric])  | 
|
| 27468 | 655  | 
|
| 64435 | 656  | 
lemma Infinitesimal_add_approx: "y \<in> Infinitesimal \<Longrightarrow> x + y = z \<Longrightarrow> x \<approx> z"  | 
657  | 
apply (rule bex_Infinitesimal_iff [THEN iffD1])  | 
|
658  | 
apply (drule Infinitesimal_minus_iff [THEN iffD2])  | 
|
659  | 
apply (auto simp add: add.assoc [symmetric])  | 
|
660  | 
done  | 
|
| 27468 | 661  | 
|
| 64435 | 662  | 
lemma Infinitesimal_add_approx_self: "y \<in> Infinitesimal \<Longrightarrow> x \<approx> x + y"  | 
663  | 
apply (rule bex_Infinitesimal_iff [THEN iffD1])  | 
|
664  | 
apply (drule Infinitesimal_minus_iff [THEN iffD2])  | 
|
665  | 
apply (auto simp add: add.assoc [symmetric])  | 
|
666  | 
done  | 
|
| 27468 | 667  | 
|
| 64435 | 668  | 
lemma Infinitesimal_add_approx_self2: "y \<in> Infinitesimal \<Longrightarrow> x \<approx> y + x"  | 
669  | 
by (auto dest: Infinitesimal_add_approx_self simp add: add.commute)  | 
|
| 27468 | 670  | 
|
| 64435 | 671  | 
lemma Infinitesimal_add_minus_approx_self: "y \<in> Infinitesimal \<Longrightarrow> x \<approx> x + - y"  | 
672  | 
by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2])  | 
|
| 27468 | 673  | 
|
| 64435 | 674  | 
lemma Infinitesimal_add_cancel: "y \<in> Infinitesimal \<Longrightarrow> x + y \<approx> z \<Longrightarrow> x \<approx> z"  | 
675  | 
apply (drule_tac x = x in Infinitesimal_add_approx_self [THEN approx_sym])  | 
|
676  | 
apply (erule approx_trans3 [THEN approx_sym], assumption)  | 
|
677  | 
done  | 
|
| 27468 | 678  | 
|
| 64435 | 679  | 
lemma Infinitesimal_add_right_cancel: "y \<in> Infinitesimal \<Longrightarrow> x \<approx> z + y \<Longrightarrow> x \<approx> z"  | 
680  | 
apply (drule_tac x = z in Infinitesimal_add_approx_self2 [THEN approx_sym])  | 
|
681  | 
apply (erule approx_trans3 [THEN approx_sym])  | 
|
682  | 
apply (simp add: add.commute)  | 
|
683  | 
apply (erule approx_sym)  | 
|
684  | 
done  | 
|
| 27468 | 685  | 
|
| 64435 | 686  | 
lemma approx_add_left_cancel: "d + b \<approx> d + c \<Longrightarrow> b \<approx> c"  | 
687  | 
apply (drule approx_minus_iff [THEN iffD1])  | 
|
688  | 
apply (simp add: approx_minus_iff [symmetric] ac_simps)  | 
|
689  | 
done  | 
|
| 27468 | 690  | 
|
| 64435 | 691  | 
lemma approx_add_right_cancel: "b + d \<approx> c + d \<Longrightarrow> b \<approx> c"  | 
692  | 
apply (rule approx_add_left_cancel)  | 
|
693  | 
apply (simp add: add.commute)  | 
|
694  | 
done  | 
|
| 27468 | 695  | 
|
| 64435 | 696  | 
lemma approx_add_mono1: "b \<approx> c \<Longrightarrow> d + b \<approx> d + c"  | 
697  | 
apply (rule approx_minus_iff [THEN iffD2])  | 
|
698  | 
apply (simp add: approx_minus_iff [symmetric] ac_simps)  | 
|
699  | 
done  | 
|
| 27468 | 700  | 
|
| 64435 | 701  | 
lemma approx_add_mono2: "b \<approx> c \<Longrightarrow> b + a \<approx> c + a"  | 
702  | 
by (simp add: add.commute approx_add_mono1)  | 
|
| 27468 | 703  | 
|
| 64435 | 704  | 
lemma approx_add_left_iff [simp]: "a + b \<approx> a + c \<longleftrightarrow> b \<approx> c"  | 
705  | 
by (fast elim: approx_add_left_cancel approx_add_mono1)  | 
|
| 27468 | 706  | 
|
| 64435 | 707  | 
lemma approx_add_right_iff [simp]: "b + a \<approx> c + a \<longleftrightarrow> b \<approx> c"  | 
708  | 
by (simp add: add.commute)  | 
|
| 27468 | 709  | 
|
| 64435 | 710  | 
lemma approx_HFinite: "x \<in> HFinite \<Longrightarrow> x \<approx> y \<Longrightarrow> y \<in> HFinite"  | 
711  | 
apply (drule bex_Infinitesimal_iff2 [THEN iffD2], safe)  | 
|
712  | 
apply (drule Infinitesimal_subset_HFinite [THEN subsetD, THEN HFinite_minus_iff [THEN iffD2]])  | 
|
713  | 
apply (drule HFinite_add)  | 
|
714  | 
apply (auto simp add: add.assoc)  | 
|
715  | 
done  | 
|
| 27468 | 716  | 
|
| 64435 | 717  | 
lemma approx_star_of_HFinite: "x \<approx> star_of D \<Longrightarrow> x \<in> HFinite"  | 
718  | 
by (rule approx_sym [THEN [2] approx_HFinite], auto)  | 
|
| 27468 | 719  | 
|
| 64435 | 720  | 
lemma approx_mult_HFinite: "a \<approx> b \<Longrightarrow> c \<approx> d \<Longrightarrow> b \<in> HFinite \<Longrightarrow> d \<in> HFinite \<Longrightarrow> a * c \<approx> b * d"  | 
721  | 
for a b c d :: "'a::real_normed_algebra star"  | 
|
722  | 
apply (rule approx_trans)  | 
|
723  | 
apply (rule_tac [2] approx_mult2)  | 
|
724  | 
apply (rule approx_mult1)  | 
|
725  | 
prefer 2 apply (blast intro: approx_HFinite approx_sym, auto)  | 
|
726  | 
done  | 
|
| 27468 | 727  | 
|
| 64435 | 728  | 
lemma scaleHR_left_diff_distrib: "\<And>a b x. scaleHR (a - b) x = scaleHR a x - scaleHR b x"  | 
729  | 
by transfer (rule scaleR_left_diff_distrib)  | 
|
| 27468 | 730  | 
|
| 64435 | 731  | 
lemma approx_scaleR1: "a \<approx> star_of b \<Longrightarrow> c \<in> HFinite \<Longrightarrow> scaleHR a c \<approx> b *\<^sub>R c"  | 
732  | 
apply (unfold approx_def)  | 
|
733  | 
apply (drule (1) Infinitesimal_HFinite_scaleHR)  | 
|
734  | 
apply (simp only: scaleHR_left_diff_distrib)  | 
|
735  | 
apply (simp add: scaleHR_def star_scaleR_def [symmetric])  | 
|
736  | 
done  | 
|
| 27468 | 737  | 
|
| 64435 | 738  | 
lemma approx_scaleR2: "a \<approx> b \<Longrightarrow> c *\<^sub>R a \<approx> c *\<^sub>R b"  | 
739  | 
by (simp add: approx_def Infinitesimal_scaleR2 scaleR_right_diff_distrib [symmetric])  | 
|
740  | 
||
741  | 
lemma approx_scaleR_HFinite: "a \<approx> star_of b \<Longrightarrow> c \<approx> d \<Longrightarrow> d \<in> HFinite \<Longrightarrow> scaleHR a c \<approx> b *\<^sub>R d"  | 
|
742  | 
apply (rule approx_trans)  | 
|
743  | 
apply (rule_tac [2] approx_scaleR2)  | 
|
744  | 
apply (rule approx_scaleR1)  | 
|
745  | 
prefer 2 apply (blast intro: approx_HFinite approx_sym, auto)  | 
|
746  | 
done  | 
|
| 27468 | 747  | 
|
| 64435 | 748  | 
lemma approx_mult_star_of: "a \<approx> star_of b \<Longrightarrow> c \<approx> star_of d \<Longrightarrow> a * c \<approx> star_of b * star_of d"  | 
749  | 
for a c :: "'a::real_normed_algebra star"  | 
|
750  | 
by (blast intro!: approx_mult_HFinite approx_star_of_HFinite HFinite_star_of)  | 
|
751  | 
||
752  | 
lemma approx_SReal_mult_cancel_zero: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a * x \<approx> 0 \<Longrightarrow> x \<approx> 0"  | 
|
753  | 
for a x :: hypreal  | 
|
754  | 
apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])  | 
|
755  | 
apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])  | 
|
756  | 
done  | 
|
| 27468 | 757  | 
|
| 64435 | 758  | 
lemma approx_mult_SReal1: "a \<in> \<real> \<Longrightarrow> x \<approx> 0 \<Longrightarrow> x * a \<approx> 0"  | 
759  | 
for a x :: hypreal  | 
|
760  | 
by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1)  | 
|
| 27468 | 761  | 
|
| 64435 | 762  | 
lemma approx_mult_SReal2: "a \<in> \<real> \<Longrightarrow> x \<approx> 0 \<Longrightarrow> a * x \<approx> 0"  | 
763  | 
for a x :: hypreal  | 
|
764  | 
by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2)  | 
|
| 27468 | 765  | 
|
| 64435 | 766  | 
lemma approx_mult_SReal_zero_cancel_iff [simp]: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a * x \<approx> 0 \<longleftrightarrow> x \<approx> 0"  | 
767  | 
for a x :: hypreal  | 
|
768  | 
by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2)  | 
|
| 27468 | 769  | 
|
| 64435 | 770  | 
lemma approx_SReal_mult_cancel: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a * w \<approx> a * z \<Longrightarrow> w \<approx> z"  | 
771  | 
for a w z :: hypreal  | 
|
772  | 
apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])  | 
|
773  | 
apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])  | 
|
774  | 
done  | 
|
| 27468 | 775  | 
|
| 64435 | 776  | 
lemma approx_SReal_mult_cancel_iff1 [simp]: "a \<in> \<real> \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a * w \<approx> a * z \<longleftrightarrow> w \<approx> z"  | 
777  | 
for a w z :: hypreal  | 
|
778  | 
by (auto intro!: approx_mult2 SReal_subset_HFinite [THEN subsetD]  | 
|
779  | 
intro: approx_SReal_mult_cancel)  | 
|
| 27468 | 780  | 
|
| 64435 | 781  | 
lemma approx_le_bound: "z \<le> f \<Longrightarrow> f \<approx> g \<Longrightarrow> g \<le> z ==> f \<approx> z"  | 
782  | 
for z :: hypreal  | 
|
783  | 
apply (simp add: bex_Infinitesimal_iff2 [symmetric], auto)  | 
|
784  | 
apply (rule_tac x = "g + y - z" in bexI)  | 
|
785  | 
apply simp  | 
|
786  | 
apply (rule Infinitesimal_interval2)  | 
|
787  | 
apply (rule_tac [2] Infinitesimal_zero, auto)  | 
|
788  | 
done  | 
|
| 27468 | 789  | 
|
| 64435 | 790  | 
lemma approx_hnorm: "x \<approx> y \<Longrightarrow> hnorm x \<approx> hnorm y"  | 
791  | 
for x y :: "'a::real_normed_vector star"  | 
|
| 27468 | 792  | 
proof (unfold approx_def)  | 
793  | 
assume "x - y \<in> Infinitesimal"  | 
|
| 64435 | 794  | 
then have "hnorm (x - y) \<in> Infinitesimal"  | 
| 27468 | 795  | 
by (simp only: Infinitesimal_hnorm_iff)  | 
| 64435 | 796  | 
moreover have "(0::real star) \<in> Infinitesimal"  | 
| 27468 | 797  | 
by (rule Infinitesimal_zero)  | 
| 64435 | 798  | 
moreover have "0 \<le> \<bar>hnorm x - hnorm y\<bar>"  | 
| 27468 | 799  | 
by (rule abs_ge_zero)  | 
| 64435 | 800  | 
moreover have "\<bar>hnorm x - hnorm y\<bar> \<le> hnorm (x - y)"  | 
| 27468 | 801  | 
by (rule hnorm_triangle_ineq3)  | 
802  | 
ultimately have "\<bar>hnorm x - hnorm y\<bar> \<in> Infinitesimal"  | 
|
803  | 
by (rule Infinitesimal_interval2)  | 
|
| 64435 | 804  | 
then show "hnorm x - hnorm y \<in> Infinitesimal"  | 
| 27468 | 805  | 
by (simp only: Infinitesimal_hrabs_iff)  | 
806  | 
qed  | 
|
807  | 
||
808  | 
||
| 64435 | 809  | 
subsection \<open>Zero is the Only Infinitesimal that is also a Real\<close>  | 
| 27468 | 810  | 
|
| 64435 | 811  | 
lemma Infinitesimal_less_SReal: "x \<in> \<real> \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> 0 < x \<Longrightarrow> y < x"  | 
812  | 
for x y :: hypreal  | 
|
813  | 
apply (simp add: Infinitesimal_def)  | 
|
814  | 
apply (rule abs_ge_self [THEN order_le_less_trans], auto)  | 
|
815  | 
done  | 
|
| 27468 | 816  | 
|
| 64435 | 817  | 
lemma Infinitesimal_less_SReal2: "y \<in> Infinitesimal \<Longrightarrow> \<forall>r \<in> Reals. 0 < r \<longrightarrow> y < r"  | 
818  | 
for y :: hypreal  | 
|
819  | 
by (blast intro: Infinitesimal_less_SReal)  | 
|
| 27468 | 820  | 
|
| 64435 | 821  | 
lemma SReal_not_Infinitesimal: "0 < y \<Longrightarrow> y \<in> \<real> ==> y \<notin> Infinitesimal"  | 
822  | 
for y :: hypreal  | 
|
823  | 
apply (simp add: Infinitesimal_def)  | 
|
824  | 
apply (auto simp add: abs_if)  | 
|
825  | 
done  | 
|
| 27468 | 826  | 
|
| 64435 | 827  | 
lemma SReal_minus_not_Infinitesimal: "y < 0 \<Longrightarrow> y \<in> \<real> \<Longrightarrow> y \<notin> Infinitesimal"  | 
828  | 
for y :: hypreal  | 
|
829  | 
apply (subst Infinitesimal_minus_iff [symmetric])  | 
|
830  | 
apply (rule SReal_not_Infinitesimal, auto)  | 
|
831  | 
done  | 
|
| 27468 | 832  | 
|
| 61070 | 833  | 
lemma SReal_Int_Infinitesimal_zero: "\<real> Int Infinitesimal = {0::hypreal}"
 | 
| 64435 | 834  | 
apply auto  | 
835  | 
apply (cut_tac x = x and y = 0 in linorder_less_linear)  | 
|
836  | 
apply (blast dest: SReal_not_Infinitesimal SReal_minus_not_Infinitesimal)  | 
|
837  | 
done  | 
|
| 27468 | 838  | 
|
| 64435 | 839  | 
lemma SReal_Infinitesimal_zero: "x \<in> \<real> \<Longrightarrow> x \<in> Infinitesimal \<Longrightarrow> x = 0"  | 
840  | 
for x :: hypreal  | 
|
841  | 
using SReal_Int_Infinitesimal_zero by blast  | 
|
| 27468 | 842  | 
|
| 64435 | 843  | 
lemma SReal_HFinite_diff_Infinitesimal: "x \<in> \<real> \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> x \<in> HFinite - Infinitesimal"  | 
844  | 
for x :: hypreal  | 
|
845  | 
by (auto dest: SReal_Infinitesimal_zero SReal_subset_HFinite [THEN subsetD])  | 
|
| 27468 | 846  | 
|
847  | 
lemma hypreal_of_real_HFinite_diff_Infinitesimal:  | 
|
| 64435 | 848  | 
"hypreal_of_real x \<noteq> 0 \<Longrightarrow> hypreal_of_real x \<in> HFinite - Infinitesimal"  | 
849  | 
by (rule SReal_HFinite_diff_Infinitesimal) auto  | 
|
| 27468 | 850  | 
|
| 64435 | 851  | 
lemma star_of_Infinitesimal_iff_0 [iff]: "star_of x \<in> Infinitesimal \<longleftrightarrow> x = 0"  | 
852  | 
apply (auto simp add: Infinitesimal_def)  | 
|
853  | 
apply (drule_tac x="hnorm (star_of x)" in bspec)  | 
|
854  | 
apply (simp add: SReal_def)  | 
|
855  | 
apply (rule_tac x="norm x" in exI, simp)  | 
|
856  | 
apply simp  | 
|
857  | 
done  | 
|
| 27468 | 858  | 
|
| 64435 | 859  | 
lemma star_of_HFinite_diff_Infinitesimal: "x \<noteq> 0 \<Longrightarrow> star_of x \<in> HFinite - Infinitesimal"  | 
860  | 
by simp  | 
|
| 27468 | 861  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
45541 
diff
changeset
 | 
862  | 
lemma numeral_not_Infinitesimal [simp]:  | 
| 64435 | 863  | 
"numeral w \<noteq> (0::hypreal) \<Longrightarrow> (numeral w :: hypreal) \<notin> Infinitesimal"  | 
864  | 
by (fast dest: Reals_numeral [THEN SReal_Infinitesimal_zero])  | 
|
| 27468 | 865  | 
|
| 64435 | 866  | 
text \<open>Again: \<open>1\<close> is a special case, but not \<open>0\<close> this time.\<close>  | 
| 27468 | 867  | 
lemma one_not_Infinitesimal [simp]:  | 
868  | 
  "(1::'a::{real_normed_vector,zero_neq_one} star) \<notin> Infinitesimal"
 | 
|
| 64435 | 869  | 
apply (simp only: star_one_def star_of_Infinitesimal_iff_0)  | 
870  | 
apply simp  | 
|
871  | 
done  | 
|
| 27468 | 872  | 
|
| 64435 | 873  | 
lemma approx_SReal_not_zero: "y \<in> \<real> \<Longrightarrow> x \<approx> y \<Longrightarrow> y \<noteq> 0 \<Longrightarrow> x \<noteq> 0"  | 
874  | 
for x y :: hypreal  | 
|
875  | 
apply (cut_tac x = 0 and y = y in linorder_less_linear, simp)  | 
|
876  | 
apply (blast dest: approx_sym [THEN mem_infmal_iff [THEN iffD2]]  | 
|
877  | 
SReal_not_Infinitesimal SReal_minus_not_Infinitesimal)  | 
|
878  | 
done  | 
|
| 27468 | 879  | 
|
880  | 
lemma HFinite_diff_Infinitesimal_approx:  | 
|
| 64435 | 881  | 
"x \<approx> y \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> x \<in> HFinite - Infinitesimal"  | 
882  | 
apply (auto intro: approx_sym [THEN [2] approx_HFinite] simp: mem_infmal_iff)  | 
|
883  | 
apply (drule approx_trans3, assumption)  | 
|
884  | 
apply (blast dest: approx_sym)  | 
|
885  | 
done  | 
|
| 27468 | 886  | 
|
| 64435 | 887  | 
text \<open>The premise \<open>y \<noteq> 0\<close> is essential; otherwise \<open>x / y = 0\<close> and we lose the  | 
888  | 
\<open>HFinite\<close> premise.\<close>  | 
|
| 27468 | 889  | 
lemma Infinitesimal_ratio:  | 
| 64435 | 890  | 
"y \<noteq> 0 \<Longrightarrow> y \<in> Infinitesimal \<Longrightarrow> x / y \<in> HFinite \<Longrightarrow> x \<in> Infinitesimal"  | 
891  | 
  for x y :: "'a::{real_normed_div_algebra,field} star"
 | 
|
892  | 
apply (drule Infinitesimal_HFinite_mult2, assumption)  | 
|
893  | 
apply (simp add: divide_inverse mult.assoc)  | 
|
894  | 
done  | 
|
895  | 
||
896  | 
lemma Infinitesimal_SReal_divide: "x \<in> Infinitesimal \<Longrightarrow> y \<in> \<real> \<Longrightarrow> x / y \<in> Infinitesimal"  | 
|
897  | 
for x y :: hypreal  | 
|
898  | 
apply (simp add: divide_inverse)  | 
|
899  | 
apply (auto intro!: Infinitesimal_HFinite_mult  | 
|
900  | 
dest!: Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]])  | 
|
901  | 
done  | 
|
902  | 
||
903  | 
||
904  | 
section \<open>Standard Part Theorem\<close>  | 
|
| 27468 | 905  | 
|
| 64435 | 906  | 
text \<open>  | 
907  | 
Every finite \<open>x \<in> R*\<close> is infinitely close to a unique real number  | 
|
908  | 
(i.e. a member of \<open>Reals\<close>).  | 
|
909  | 
\<close>  | 
|
| 27468 | 910  | 
|
911  | 
||
| 64435 | 912  | 
subsection \<open>Uniqueness: Two Infinitely Close Reals are Equal\<close>  | 
| 27468 | 913  | 
|
| 64435 | 914  | 
lemma star_of_approx_iff [simp]: "star_of x \<approx> star_of y \<longleftrightarrow> x = y"  | 
915  | 
apply safe  | 
|
916  | 
apply (simp add: approx_def)  | 
|
917  | 
apply (simp only: star_of_diff [symmetric])  | 
|
918  | 
apply (simp only: star_of_Infinitesimal_iff_0)  | 
|
919  | 
apply simp  | 
|
920  | 
done  | 
|
| 27468 | 921  | 
|
| 64435 | 922  | 
lemma SReal_approx_iff: "x \<in> \<real> \<Longrightarrow> y \<in> \<real> \<Longrightarrow> x \<approx> y \<longleftrightarrow> x = y"  | 
923  | 
for x y :: hypreal  | 
|
924  | 
apply auto  | 
|
925  | 
apply (simp add: approx_def)  | 
|
926  | 
apply (drule (1) Reals_diff)  | 
|
927  | 
apply (drule (1) SReal_Infinitesimal_zero)  | 
|
928  | 
apply simp  | 
|
929  | 
done  | 
|
| 27468 | 930  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
45541 
diff
changeset
 | 
931  | 
lemma numeral_approx_iff [simp]:  | 
| 64435 | 932  | 
  "(numeral v \<approx> (numeral w :: 'a::{numeral,real_normed_vector} star)) =
 | 
933  | 
(numeral v = (numeral w :: 'a))"  | 
|
934  | 
apply (unfold star_numeral_def)  | 
|
935  | 
apply (rule star_of_approx_iff)  | 
|
936  | 
done  | 
|
| 27468 | 937  | 
|
| 64435 | 938  | 
text \<open>And also for \<open>0 \<approx> #nn\<close> and \<open>1 \<approx> #nn\<close>, \<open>#nn \<approx> 0\<close> and \<open>#nn \<approx> 1\<close>.\<close>  | 
| 27468 | 939  | 
lemma [simp]:  | 
| 64435 | 940  | 
  "(numeral w \<approx> (0::'a::{numeral,real_normed_vector} star)) = (numeral w = (0::'a))"
 | 
941  | 
  "((0::'a::{numeral,real_normed_vector} star) \<approx> numeral w) = (numeral w = (0::'a))"
 | 
|
942  | 
  "(numeral w \<approx> (1::'b::{numeral,one,real_normed_vector} star)) = (numeral w = (1::'b))"
 | 
|
943  | 
  "((1::'b::{numeral,one,real_normed_vector} star) \<approx> numeral w) = (numeral w = (1::'b))"
 | 
|
944  | 
  "\<not> (0 \<approx> (1::'c::{zero_neq_one,real_normed_vector} star))"
 | 
|
945  | 
  "\<not> (1 \<approx> (0::'c::{zero_neq_one,real_normed_vector} star))"
 | 
|
946  | 
apply (unfold star_numeral_def star_zero_def star_one_def)  | 
|
947  | 
apply (unfold star_of_approx_iff)  | 
|
948  | 
apply (auto intro: sym)  | 
|
949  | 
done  | 
|
| 27468 | 950  | 
|
| 64435 | 951  | 
lemma star_of_approx_numeral_iff [simp]: "star_of k \<approx> numeral w \<longleftrightarrow> k = numeral w"  | 
952  | 
by (subst star_of_approx_iff [symmetric]) auto  | 
|
| 27468 | 953  | 
|
| 64435 | 954  | 
lemma star_of_approx_zero_iff [simp]: "star_of k \<approx> 0 \<longleftrightarrow> k = 0"  | 
955  | 
by (simp_all add: star_of_approx_iff [symmetric])  | 
|
| 27468 | 956  | 
|
| 64435 | 957  | 
lemma star_of_approx_one_iff [simp]: "star_of k \<approx> 1 \<longleftrightarrow> k = 1"  | 
958  | 
by (simp_all add: star_of_approx_iff [symmetric])  | 
|
| 27468 | 959  | 
|
| 64435 | 960  | 
lemma approx_unique_real: "r \<in> \<real> \<Longrightarrow> s \<in> \<real> \<Longrightarrow> r \<approx> x \<Longrightarrow> s \<approx> x \<Longrightarrow> r = s"  | 
961  | 
for r s :: hypreal  | 
|
962  | 
by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2)  | 
|
| 27468 | 963  | 
|
964  | 
||
| 64435 | 965  | 
subsection \<open>Existence of Unique Real Infinitely Close\<close>  | 
| 27468 | 966  | 
|
| 64435 | 967  | 
subsubsection \<open>Lifting of the Ub and Lub Properties\<close>  | 
| 27468 | 968  | 
|
| 64435 | 969  | 
lemma hypreal_of_real_isUb_iff: "isUb \<real> (hypreal_of_real ` Q) (hypreal_of_real Y) = isUb UNIV Q Y"  | 
970  | 
for Q :: "real set" and Y :: real  | 
|
971  | 
by (simp add: isUb_def setle_def)  | 
|
| 27468 | 972  | 
|
| 64435 | 973  | 
lemma hypreal_of_real_isLub1: "isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y) \<Longrightarrow> isLub UNIV Q Y"  | 
974  | 
for Q :: "real set" and Y :: real  | 
|
975  | 
apply (simp add: isLub_def leastP_def)  | 
|
976  | 
apply (auto intro: hypreal_of_real_isUb_iff [THEN iffD2]  | 
|
977  | 
simp add: hypreal_of_real_isUb_iff setge_def)  | 
|
978  | 
done  | 
|
| 27468 | 979  | 
|
| 64435 | 980  | 
lemma hypreal_of_real_isLub2: "isLub UNIV Q Y \<Longrightarrow> isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y)"  | 
981  | 
for Q :: "real set" and Y :: real  | 
|
982  | 
apply (auto simp add: isLub_def leastP_def hypreal_of_real_isUb_iff setge_def)  | 
|
983  | 
apply (metis SReal_iff hypreal_of_real_isUb_iff isUbD2a star_of_le)  | 
|
984  | 
done  | 
|
| 27468 | 985  | 
|
986  | 
lemma hypreal_of_real_isLub_iff:  | 
|
| 64435 | 987  | 
"isLub \<real> (hypreal_of_real ` Q) (hypreal_of_real Y) = isLub (UNIV :: real set) Q Y"  | 
988  | 
for Q :: "real set" and Y :: real  | 
|
989  | 
by (blast intro: hypreal_of_real_isLub1 hypreal_of_real_isLub2)  | 
|
| 27468 | 990  | 
|
| 64435 | 991  | 
lemma lemma_isUb_hypreal_of_real: "isUb \<real> P Y \<Longrightarrow> \<exists>Yo. isUb \<real> P (hypreal_of_real Yo)"  | 
992  | 
by (auto simp add: SReal_iff isUb_def)  | 
|
993  | 
||
994  | 
lemma lemma_isLub_hypreal_of_real: "isLub \<real> P Y \<Longrightarrow> \<exists>Yo. isLub \<real> P (hypreal_of_real Yo)"  | 
|
995  | 
by (auto simp add: isLub_def leastP_def isUb_def SReal_iff)  | 
|
| 27468 | 996  | 
|
| 64435 | 997  | 
lemma lemma_isLub_hypreal_of_real2: "\<exists>Yo. isLub \<real> P (hypreal_of_real Yo) \<Longrightarrow> \<exists>Y. isLub \<real> P Y"  | 
998  | 
by (auto simp add: isLub_def leastP_def isUb_def)  | 
|
| 27468 | 999  | 
|
| 64435 | 1000  | 
lemma SReal_complete: "P \<subseteq> \<real> \<Longrightarrow> \<exists>x. x \<in> P \<Longrightarrow> \<exists>Y. isUb \<real> P Y \<Longrightarrow> \<exists>t::hypreal. isLub \<real> P t"  | 
1001  | 
apply (frule SReal_hypreal_of_real_image)  | 
|
1002  | 
apply (auto, drule lemma_isUb_hypreal_of_real)  | 
|
1003  | 
apply (auto intro!: reals_complete lemma_isLub_hypreal_of_real2  | 
|
1004  | 
simp add: hypreal_of_real_isLub_iff hypreal_of_real_isUb_iff)  | 
|
1005  | 
done  | 
|
1006  | 
||
| 27468 | 1007  | 
|
| 64435 | 1008  | 
text \<open>Lemmas about lubs.\<close>  | 
| 27468 | 1009  | 
|
| 64435 | 1010  | 
lemma lemma_st_part_ub: "x \<in> HFinite \<Longrightarrow> \<exists>u. isUb \<real> {s. s \<in> \<real> \<and> s < x} u"
 | 
1011  | 
for x :: hypreal  | 
|
1012  | 
apply (drule HFiniteD, safe)  | 
|
1013  | 
apply (rule exI, rule isUbI)  | 
|
1014  | 
apply (auto intro: setleI isUbI simp add: abs_less_iff)  | 
|
1015  | 
done  | 
|
| 27468 | 1016  | 
|
| 64435 | 1017  | 
lemma lemma_st_part_nonempty: "x \<in> HFinite \<Longrightarrow> \<exists>y. y \<in> {s. s \<in> \<real> \<and> s < x}"
 | 
1018  | 
for x :: hypreal  | 
|
1019  | 
apply (drule HFiniteD, safe)  | 
|
1020  | 
apply (drule Reals_minus)  | 
|
1021  | 
apply (rule_tac x = "-t" in exI)  | 
|
1022  | 
apply (auto simp add: abs_less_iff)  | 
|
1023  | 
done  | 
|
| 27468 | 1024  | 
|
| 64435 | 1025  | 
lemma lemma_st_part_lub: "x \<in> HFinite \<Longrightarrow> \<exists>t. isLub \<real> {s. s \<in> \<real> \<and> s < x} t"
 | 
1026  | 
for x :: hypreal  | 
|
1027  | 
by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty Collect_restrict)  | 
|
| 27468 | 1028  | 
|
1029  | 
lemma lemma_st_part_le1:  | 
|
| 64435 | 1030  | 
  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> x \<le> t + r"
 | 
1031  | 
for x r t :: hypreal  | 
|
1032  | 
apply (frule isLubD1a)  | 
|
1033  | 
apply (rule ccontr, drule linorder_not_le [THEN iffD2])  | 
|
1034  | 
apply (drule (1) Reals_add)  | 
|
1035  | 
apply (drule_tac y = "r + t" in isLubD1 [THEN setleD], auto)  | 
|
1036  | 
done  | 
|
| 27468 | 1037  | 
|
| 64435 | 1038  | 
lemma hypreal_setle_less_trans: "S *<= x \<Longrightarrow> x < y \<Longrightarrow> S *<= y"  | 
1039  | 
for x y :: hypreal  | 
|
1040  | 
apply (simp add: setle_def)  | 
|
1041  | 
apply (auto dest!: bspec order_le_less_trans intro: order_less_imp_le)  | 
|
1042  | 
done  | 
|
| 27468 | 1043  | 
|
| 64435 | 1044  | 
lemma hypreal_gt_isUb: "isUb R S x \<Longrightarrow> x < y \<Longrightarrow> y \<in> R \<Longrightarrow> isUb R S y"  | 
1045  | 
for x y :: hypreal  | 
|
1046  | 
apply (simp add: isUb_def)  | 
|
1047  | 
apply (blast intro: hypreal_setle_less_trans)  | 
|
1048  | 
done  | 
|
| 27468 | 1049  | 
|
| 64435 | 1050  | 
lemma lemma_st_part_gt_ub: "x \<in> HFinite \<Longrightarrow> x < y \<Longrightarrow> y \<in> \<real> \<Longrightarrow> isUb \<real> {s. s \<in> \<real> \<and> s < x} y"
 | 
1051  | 
for x y :: hypreal  | 
|
1052  | 
by (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI)  | 
|
| 27468 | 1053  | 
|
| 64435 | 1054  | 
lemma lemma_minus_le_zero: "t \<le> t + -r \<Longrightarrow> r \<le> 0"  | 
1055  | 
for r t :: hypreal  | 
|
1056  | 
apply (drule_tac c = "-t" in add_left_mono)  | 
|
1057  | 
apply (auto simp add: add.assoc [symmetric])  | 
|
1058  | 
done  | 
|
| 27468 | 1059  | 
|
1060  | 
lemma lemma_st_part_le2:  | 
|
| 64435 | 1061  | 
  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> t + -r \<le> x"
 | 
1062  | 
for x r t :: hypreal  | 
|
1063  | 
apply (frule isLubD1a)  | 
|
1064  | 
apply (rule ccontr, drule linorder_not_le [THEN iffD1])  | 
|
1065  | 
apply (drule Reals_minus, drule_tac a = t in Reals_add, assumption)  | 
|
1066  | 
apply (drule lemma_st_part_gt_ub, assumption+)  | 
|
1067  | 
apply (drule isLub_le_isUb, assumption)  | 
|
1068  | 
apply (drule lemma_minus_le_zero)  | 
|
1069  | 
apply (auto dest: order_less_le_trans)  | 
|
1070  | 
done  | 
|
| 27468 | 1071  | 
|
1072  | 
lemma lemma_st_part1a:  | 
|
| 64435 | 1073  | 
  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> x + -t \<le> r"
 | 
1074  | 
for x r t :: hypreal  | 
|
1075  | 
apply (subgoal_tac "x \<le> t + r")  | 
|
1076  | 
apply (auto intro: lemma_st_part_le1)  | 
|
1077  | 
done  | 
|
| 27468 | 1078  | 
|
1079  | 
lemma lemma_st_part2a:  | 
|
| 64435 | 1080  | 
  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> - (x + -t) \<le> r"
 | 
1081  | 
for x r t :: hypreal  | 
|
1082  | 
apply (subgoal_tac "(t + -r \<le> x)")  | 
|
1083  | 
apply simp  | 
|
1084  | 
apply (rule lemma_st_part_le2)  | 
|
1085  | 
apply auto  | 
|
1086  | 
done  | 
|
| 27468 | 1087  | 
|
| 64435 | 1088  | 
lemma lemma_SReal_ub: "x \<in> \<real> \<Longrightarrow> isUb \<real> {s. s \<in> \<real> \<and> s < x} x"
 | 
1089  | 
for x :: hypreal  | 
|
1090  | 
by (auto intro: isUbI setleI order_less_imp_le)  | 
|
| 27468 | 1091  | 
|
| 64435 | 1092  | 
lemma lemma_SReal_lub: "x \<in> \<real> \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} x"
 | 
1093  | 
for x :: hypreal  | 
|
1094  | 
apply (auto intro!: isLubI2 lemma_SReal_ub setgeI)  | 
|
1095  | 
apply (frule isUbD2a)  | 
|
1096  | 
apply (rule_tac x = x and y = y in linorder_cases)  | 
|
1097  | 
apply (auto intro!: order_less_imp_le)  | 
|
1098  | 
apply (drule SReal_dense, assumption, assumption, safe)  | 
|
1099  | 
apply (drule_tac y = r in isUbD)  | 
|
1100  | 
apply (auto dest: order_less_le_trans)  | 
|
1101  | 
done  | 
|
| 27468 | 1102  | 
|
1103  | 
lemma lemma_st_part_not_eq1:  | 
|
| 64435 | 1104  | 
  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> x + - t \<noteq> r"
 | 
1105  | 
for x r t :: hypreal  | 
|
1106  | 
apply auto  | 
|
1107  | 
apply (frule isLubD1a [THEN Reals_minus])  | 
|
1108  | 
using Reals_add_cancel [of x "- t"] apply simp  | 
|
1109  | 
apply (drule_tac x = x in lemma_SReal_lub)  | 
|
1110  | 
apply (drule isLub_unique, assumption, auto)  | 
|
1111  | 
done  | 
|
| 27468 | 1112  | 
|
1113  | 
lemma lemma_st_part_not_eq2:  | 
|
| 64435 | 1114  | 
  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> - (x + -t) \<noteq> r"
 | 
1115  | 
for x r t :: hypreal  | 
|
1116  | 
apply (auto)  | 
|
1117  | 
apply (frule isLubD1a)  | 
|
1118  | 
using Reals_add_cancel [of "- x" t] apply simp  | 
|
1119  | 
apply (drule_tac x = x in lemma_SReal_lub)  | 
|
1120  | 
apply (drule isLub_unique, assumption, auto)  | 
|
1121  | 
done  | 
|
| 27468 | 1122  | 
|
1123  | 
lemma lemma_st_part_major:  | 
|
| 64435 | 1124  | 
  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> r \<in> \<real> \<Longrightarrow> 0 < r \<Longrightarrow> \<bar>x - t\<bar> < r"
 | 
1125  | 
for x r t :: hypreal  | 
|
1126  | 
apply (frule lemma_st_part1a)  | 
|
1127  | 
apply (frule_tac [4] lemma_st_part2a, auto)  | 
|
1128  | 
apply (drule order_le_imp_less_or_eq)+  | 
|
1129  | 
apply auto  | 
|
1130  | 
using lemma_st_part_not_eq2 apply fastforce  | 
|
1131  | 
using lemma_st_part_not_eq1 apply fastforce  | 
|
1132  | 
done  | 
|
| 27468 | 1133  | 
|
1134  | 
lemma lemma_st_part_major2:  | 
|
| 64435 | 1135  | 
  "x \<in> HFinite \<Longrightarrow> isLub \<real> {s. s \<in> \<real> \<and> s < x} t \<Longrightarrow> \<forall>r \<in> Reals. 0 < r \<longrightarrow> \<bar>x - t\<bar> < r"
 | 
1136  | 
for x t :: hypreal  | 
|
1137  | 
by (blast dest!: lemma_st_part_major)  | 
|
| 27468 | 1138  | 
|
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61609 
diff
changeset
 | 
1139  | 
|
| 64435 | 1140  | 
text\<open>Existence of real and Standard Part Theorem.\<close>  | 
1141  | 
||
1142  | 
lemma lemma_st_part_Ex: "x \<in> HFinite \<Longrightarrow> \<exists>t\<in>Reals. \<forall>r \<in> Reals. 0 < r \<longrightarrow> \<bar>x - t\<bar> < r"  | 
|
1143  | 
for x :: hypreal  | 
|
1144  | 
apply (frule lemma_st_part_lub, safe)  | 
|
1145  | 
apply (frule isLubD1a)  | 
|
1146  | 
apply (blast dest: lemma_st_part_major2)  | 
|
1147  | 
done  | 
|
| 27468 | 1148  | 
|
| 64435 | 1149  | 
lemma st_part_Ex: "x \<in> HFinite \<Longrightarrow> \<exists>t\<in>Reals. x \<approx> t"  | 
1150  | 
for x :: hypreal  | 
|
1151  | 
apply (simp add: approx_def Infinitesimal_def)  | 
|
1152  | 
apply (drule lemma_st_part_Ex, auto)  | 
|
1153  | 
done  | 
|
| 27468 | 1154  | 
|
| 64435 | 1155  | 
text \<open>There is a unique real infinitely close.\<close>  | 
1156  | 
lemma st_part_Ex1: "x \<in> HFinite \<Longrightarrow> \<exists>!t::hypreal. t \<in> \<real> \<and> x \<approx> t"  | 
|
1157  | 
apply (drule st_part_Ex, safe)  | 
|
1158  | 
apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym)  | 
|
1159  | 
apply (auto intro!: approx_unique_real)  | 
|
1160  | 
done  | 
|
| 27468 | 1161  | 
|
| 64435 | 1162  | 
|
1163  | 
subsection \<open>Finite, Infinite and Infinitesimal\<close>  | 
|
| 27468 | 1164  | 
|
1165  | 
lemma HFinite_Int_HInfinite_empty [simp]: "HFinite Int HInfinite = {}"
 | 
|
| 64435 | 1166  | 
apply (simp add: HFinite_def HInfinite_def)  | 
1167  | 
apply (auto dest: order_less_trans)  | 
|
1168  | 
done  | 
|
| 27468 | 1169  | 
|
| 
56217
 
dc429a5b13c4
Some rationalisation of basic lemmas
 
paulson <lp15@cam.ac.uk> 
parents: 
54489 
diff
changeset
 | 
1170  | 
lemma HFinite_not_HInfinite:  | 
| 64435 | 1171  | 
assumes x: "x \<in> HFinite"  | 
1172  | 
shows "x \<notin> HInfinite"  | 
|
| 27468 | 1173  | 
proof  | 
1174  | 
assume x': "x \<in> HInfinite"  | 
|
1175  | 
with x have "x \<in> HFinite \<inter> HInfinite" by blast  | 
|
| 64435 | 1176  | 
then show False by auto  | 
| 27468 | 1177  | 
qed  | 
1178  | 
||
| 64435 | 1179  | 
lemma not_HFinite_HInfinite: "x \<notin> HFinite \<Longrightarrow> x \<in> HInfinite"  | 
1180  | 
apply (simp add: HInfinite_def HFinite_def, auto)  | 
|
1181  | 
apply (drule_tac x = "r + 1" in bspec)  | 
|
1182  | 
apply (auto)  | 
|
1183  | 
done  | 
|
| 27468 | 1184  | 
|
| 64435 | 1185  | 
lemma HInfinite_HFinite_disj: "x \<in> HInfinite \<or> x \<in> HFinite"  | 
1186  | 
by (blast intro: not_HFinite_HInfinite)  | 
|
| 27468 | 1187  | 
|
| 64435 | 1188  | 
lemma HInfinite_HFinite_iff: "x \<in> HInfinite \<longleftrightarrow> x \<notin> HFinite"  | 
1189  | 
by (blast dest: HFinite_not_HInfinite not_HFinite_HInfinite)  | 
|
| 27468 | 1190  | 
|
| 64435 | 1191  | 
lemma HFinite_HInfinite_iff: "x \<in> HFinite \<longleftrightarrow> x \<notin> HInfinite"  | 
1192  | 
by (simp add: HInfinite_HFinite_iff)  | 
|
| 27468 | 1193  | 
|
1194  | 
||
1195  | 
lemma HInfinite_diff_HFinite_Infinitesimal_disj:  | 
|
| 64435 | 1196  | 
"x \<notin> Infinitesimal \<Longrightarrow> x \<in> HInfinite \<or> x \<in> HFinite - Infinitesimal"  | 
1197  | 
by (fast intro: not_HFinite_HInfinite)  | 
|
| 27468 | 1198  | 
|
| 64435 | 1199  | 
lemma HFinite_inverse: "x \<in> HFinite \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> inverse x \<in> HFinite"  | 
1200  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
1201  | 
apply (subgoal_tac "x \<noteq> 0")  | 
|
1202  | 
apply (cut_tac x = "inverse x" in HInfinite_HFinite_disj)  | 
|
1203  | 
apply (auto dest!: HInfinite_inverse_Infinitesimal simp: nonzero_inverse_inverse_eq)  | 
|
1204  | 
done  | 
|
| 27468 | 1205  | 
|
| 64435 | 1206  | 
lemma HFinite_inverse2: "x \<in> HFinite - Infinitesimal \<Longrightarrow> inverse x \<in> HFinite"  | 
1207  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
1208  | 
by (blast intro: HFinite_inverse)  | 
|
| 27468 | 1209  | 
|
| 64435 | 1210  | 
text \<open>Stronger statement possible in fact.\<close>  | 
1211  | 
lemma Infinitesimal_inverse_HFinite: "x \<notin> Infinitesimal \<Longrightarrow> inverse x \<in> HFinite"  | 
|
1212  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
1213  | 
apply (drule HInfinite_diff_HFinite_Infinitesimal_disj)  | 
|
1214  | 
apply (blast intro: HFinite_inverse HInfinite_inverse_Infinitesimal Infinitesimal_subset_HFinite [THEN subsetD])  | 
|
1215  | 
done  | 
|
| 27468 | 1216  | 
|
1217  | 
lemma HFinite_not_Infinitesimal_inverse:  | 
|
| 64435 | 1218  | 
"x \<in> HFinite - Infinitesimal \<Longrightarrow> inverse x \<in> HFinite - Infinitesimal"  | 
1219  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
1220  | 
apply (auto intro: Infinitesimal_inverse_HFinite)  | 
|
1221  | 
apply (drule Infinitesimal_HFinite_mult2, assumption)  | 
|
1222  | 
apply (simp add: not_Infinitesimal_not_zero)  | 
|
1223  | 
done  | 
|
| 27468 | 1224  | 
|
| 64435 | 1225  | 
lemma approx_inverse: "x \<approx> y \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> inverse x \<approx> inverse y"  | 
1226  | 
for x y :: "'a::real_normed_div_algebra star"  | 
|
1227  | 
apply (frule HFinite_diff_Infinitesimal_approx, assumption)  | 
|
1228  | 
apply (frule not_Infinitesimal_not_zero2)  | 
|
1229  | 
apply (frule_tac x = x in not_Infinitesimal_not_zero2)  | 
|
1230  | 
apply (drule HFinite_inverse2)+  | 
|
1231  | 
apply (drule approx_mult2, assumption, auto)  | 
|
1232  | 
apply (drule_tac c = "inverse x" in approx_mult1, assumption)  | 
|
1233  | 
apply (auto intro: approx_sym simp add: mult.assoc)  | 
|
1234  | 
done  | 
|
| 27468 | 1235  | 
|
1236  | 
(*Used for NSLIM_inverse, NSLIMSEQ_inverse*)  | 
|
1237  | 
lemmas star_of_approx_inverse = star_of_HFinite_diff_Infinitesimal [THEN [2] approx_inverse]  | 
|
1238  | 
lemmas hypreal_of_real_approx_inverse = hypreal_of_real_HFinite_diff_Infinitesimal [THEN [2] approx_inverse]  | 
|
1239  | 
||
1240  | 
lemma inverse_add_Infinitesimal_approx:  | 
|
| 64435 | 1241  | 
"x \<in> HFinite - Infinitesimal \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> inverse (x + h) \<approx> inverse x"  | 
1242  | 
for x h :: "'a::real_normed_div_algebra star"  | 
|
1243  | 
by (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self)  | 
|
| 27468 | 1244  | 
|
1245  | 
lemma inverse_add_Infinitesimal_approx2:  | 
|
| 64435 | 1246  | 
"x \<in> HFinite - Infinitesimal \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> inverse (h + x) \<approx> inverse x"  | 
1247  | 
for x h :: "'a::real_normed_div_algebra star"  | 
|
1248  | 
apply (rule add.commute [THEN subst])  | 
|
1249  | 
apply (blast intro: inverse_add_Infinitesimal_approx)  | 
|
1250  | 
done  | 
|
| 27468 | 1251  | 
|
1252  | 
lemma inverse_add_Infinitesimal_approx_Infinitesimal:  | 
|
| 64435 | 1253  | 
"x \<in> HFinite - Infinitesimal \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> inverse (x + h) - inverse x \<approx> h"  | 
1254  | 
for x h :: "'a::real_normed_div_algebra star"  | 
|
1255  | 
apply (rule approx_trans2)  | 
|
1256  | 
apply (auto intro: inverse_add_Infinitesimal_approx  | 
|
1257  | 
simp add: mem_infmal_iff approx_minus_iff [symmetric])  | 
|
1258  | 
done  | 
|
| 27468 | 1259  | 
|
| 64435 | 1260  | 
lemma Infinitesimal_square_iff: "x \<in> Infinitesimal \<longleftrightarrow> x * x \<in> Infinitesimal"  | 
1261  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
1262  | 
apply (auto intro: Infinitesimal_mult)  | 
|
1263  | 
apply (rule ccontr, frule Infinitesimal_inverse_HFinite)  | 
|
1264  | 
apply (frule not_Infinitesimal_not_zero)  | 
|
1265  | 
apply (auto dest: Infinitesimal_HFinite_mult simp add: mult.assoc)  | 
|
1266  | 
done  | 
|
| 27468 | 1267  | 
declare Infinitesimal_square_iff [symmetric, simp]  | 
1268  | 
||
| 64435 | 1269  | 
lemma HFinite_square_iff [simp]: "x * x \<in> HFinite \<longleftrightarrow> x \<in> HFinite"  | 
1270  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
1271  | 
apply (auto intro: HFinite_mult)  | 
|
1272  | 
apply (auto dest: HInfinite_mult simp add: HFinite_HInfinite_iff)  | 
|
1273  | 
done  | 
|
| 27468 | 1274  | 
|
| 64435 | 1275  | 
lemma HInfinite_square_iff [simp]: "x * x \<in> HInfinite \<longleftrightarrow> x \<in> HInfinite"  | 
1276  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
1277  | 
by (auto simp add: HInfinite_HFinite_iff)  | 
|
| 27468 | 1278  | 
|
| 64435 | 1279  | 
lemma approx_HFinite_mult_cancel: "a \<in> HFinite - Infinitesimal \<Longrightarrow> a * w \<approx> a * z \<Longrightarrow> w \<approx> z"  | 
1280  | 
for a w z :: "'a::real_normed_div_algebra star"  | 
|
1281  | 
apply safe  | 
|
1282  | 
apply (frule HFinite_inverse, assumption)  | 
|
1283  | 
apply (drule not_Infinitesimal_not_zero)  | 
|
1284  | 
apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric])  | 
|
1285  | 
done  | 
|
| 27468 | 1286  | 
|
| 64435 | 1287  | 
lemma approx_HFinite_mult_cancel_iff1: "a \<in> HFinite - Infinitesimal \<Longrightarrow> a * w \<approx> a * z \<longleftrightarrow> w \<approx> z"  | 
1288  | 
for a w z :: "'a::real_normed_div_algebra star"  | 
|
1289  | 
by (auto intro: approx_mult2 approx_HFinite_mult_cancel)  | 
|
| 27468 | 1290  | 
|
| 64435 | 1291  | 
lemma HInfinite_HFinite_add_cancel: "x + y \<in> HInfinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x \<in> HInfinite"  | 
1292  | 
apply (rule ccontr)  | 
|
1293  | 
apply (drule HFinite_HInfinite_iff [THEN iffD2])  | 
|
1294  | 
apply (auto dest: HFinite_add simp add: HInfinite_HFinite_iff)  | 
|
1295  | 
done  | 
|
| 27468 | 1296  | 
|
| 64435 | 1297  | 
lemma HInfinite_HFinite_add: "x \<in> HInfinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x + y \<in> HInfinite"  | 
1298  | 
apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel)  | 
|
1299  | 
apply (auto simp add: add.assoc HFinite_minus_iff)  | 
|
1300  | 
done  | 
|
| 27468 | 1301  | 
|
| 64435 | 1302  | 
lemma HInfinite_ge_HInfinite: "x \<in> HInfinite \<Longrightarrow> x \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> y \<in> HInfinite"  | 
1303  | 
for x y :: hypreal  | 
|
1304  | 
by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff)  | 
|
| 27468 | 1305  | 
|
| 64435 | 1306  | 
lemma Infinitesimal_inverse_HInfinite: "x \<in> Infinitesimal \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> inverse x \<in> HInfinite"  | 
1307  | 
for x :: "'a::real_normed_div_algebra star"  | 
|
1308  | 
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])  | 
|
1309  | 
apply (auto dest: Infinitesimal_HFinite_mult2)  | 
|
1310  | 
done  | 
|
| 27468 | 1311  | 
|
1312  | 
lemma HInfinite_HFinite_not_Infinitesimal_mult:  | 
|
| 64435 | 1313  | 
"x \<in> HInfinite \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> x * y \<in> HInfinite"  | 
1314  | 
for x y :: "'a::real_normed_div_algebra star"  | 
|
1315  | 
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])  | 
|
1316  | 
apply (frule HFinite_Infinitesimal_not_zero)  | 
|
1317  | 
apply (drule HFinite_not_Infinitesimal_inverse)  | 
|
1318  | 
apply (safe, drule HFinite_mult)  | 
|
1319  | 
apply (auto simp add: mult.assoc HFinite_HInfinite_iff)  | 
|
1320  | 
done  | 
|
| 27468 | 1321  | 
|
1322  | 
lemma HInfinite_HFinite_not_Infinitesimal_mult2:  | 
|
| 64435 | 1323  | 
"x \<in> HInfinite \<Longrightarrow> y \<in> HFinite - Infinitesimal \<Longrightarrow> y * x \<in> HInfinite"  | 
1324  | 
for x y :: "'a::real_normed_div_algebra star"  | 
|
1325  | 
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])  | 
|
1326  | 
apply (frule HFinite_Infinitesimal_not_zero)  | 
|
1327  | 
apply (drule HFinite_not_Infinitesimal_inverse)  | 
|
1328  | 
apply (safe, drule_tac x="inverse y" in HFinite_mult)  | 
|
1329  | 
apply assumption  | 
|
1330  | 
apply (auto simp add: mult.assoc [symmetric] HFinite_HInfinite_iff)  | 
|
1331  | 
done  | 
|
| 27468 | 1332  | 
|
| 64435 | 1333  | 
lemma HInfinite_gt_SReal: "x \<in> HInfinite \<Longrightarrow> 0 < x \<Longrightarrow> y \<in> \<real> \<Longrightarrow> y < x"  | 
1334  | 
for x y :: hypreal  | 
|
1335  | 
by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le)  | 
|
| 27468 | 1336  | 
|
| 64435 | 1337  | 
lemma HInfinite_gt_zero_gt_one: "x \<in> HInfinite \<Longrightarrow> 0 < x \<Longrightarrow> 1 < x"  | 
1338  | 
for x :: hypreal  | 
|
1339  | 
by (auto intro: HInfinite_gt_SReal)  | 
|
| 27468 | 1340  | 
|
1341  | 
||
1342  | 
lemma not_HInfinite_one [simp]: "1 \<notin> HInfinite"  | 
|
| 64435 | 1343  | 
by (simp add: HInfinite_HFinite_iff)  | 
| 27468 | 1344  | 
|
| 64435 | 1345  | 
lemma approx_hrabs_disj: "\<bar>x\<bar> \<approx> x \<or> \<bar>x\<bar> \<approx> -x"  | 
1346  | 
for x :: hypreal  | 
|
1347  | 
using hrabs_disj [of x] by auto  | 
|
| 27468 | 1348  | 
|
1349  | 
||
| 64435 | 1350  | 
subsection \<open>Theorems about Monads\<close>  | 
| 27468 | 1351  | 
|
| 64435 | 1352  | 
lemma monad_hrabs_Un_subset: "monad \<bar>x\<bar> \<le> monad x \<union> monad (- x)"  | 
1353  | 
for x :: hypreal  | 
|
1354  | 
by (rule hrabs_disj [of x, THEN disjE]) auto  | 
|
| 27468 | 1355  | 
|
| 64435 | 1356  | 
lemma Infinitesimal_monad_eq: "e \<in> Infinitesimal \<Longrightarrow> monad (x + e) = monad x"  | 
1357  | 
by (fast intro!: Infinitesimal_add_approx_self [THEN approx_sym] approx_monad_iff [THEN iffD1])  | 
|
| 27468 | 1358  | 
|
| 64435 | 1359  | 
lemma mem_monad_iff: "u \<in> monad x \<longleftrightarrow> - u \<in> monad (- x)"  | 
1360  | 
by (simp add: monad_def)  | 
|
1361  | 
||
1362  | 
lemma Infinitesimal_monad_zero_iff: "x \<in> Infinitesimal \<longleftrightarrow> x \<in> monad 0"  | 
|
1363  | 
by (auto intro: approx_sym simp add: monad_def mem_infmal_iff)  | 
|
| 27468 | 1364  | 
|
| 64435 | 1365  | 
lemma monad_zero_minus_iff: "x \<in> monad 0 \<longleftrightarrow> - x \<in> monad 0"  | 
1366  | 
by (simp add: Infinitesimal_monad_zero_iff [symmetric])  | 
|
| 27468 | 1367  | 
|
| 64435 | 1368  | 
lemma monad_zero_hrabs_iff: "x \<in> monad 0 \<longleftrightarrow> \<bar>x\<bar> \<in> monad 0"  | 
1369  | 
for x :: hypreal  | 
|
1370  | 
by (rule hrabs_disj [of x, THEN disjE]) (auto simp: monad_zero_minus_iff [symmetric])  | 
|
| 27468 | 1371  | 
|
1372  | 
lemma mem_monad_self [simp]: "x \<in> monad x"  | 
|
| 64435 | 1373  | 
by (simp add: monad_def)  | 
| 27468 | 1374  | 
|
1375  | 
||
| 64435 | 1376  | 
subsection \<open>Proof that @{term "x \<approx> y"} implies @{term"\<bar>x\<bar> \<approx> \<bar>y\<bar>"}\<close>
 | 
| 27468 | 1377  | 
|
| 64435 | 1378  | 
lemma approx_subset_monad: "x \<approx> y \<Longrightarrow> {x, y} \<le> monad x"
 | 
1379  | 
by (simp (no_asm)) (simp add: approx_monad_iff)  | 
|
| 27468 | 1380  | 
|
| 64435 | 1381  | 
lemma approx_subset_monad2: "x \<approx> y \<Longrightarrow> {x, y} \<le> monad y"
 | 
1382  | 
apply (drule approx_sym)  | 
|
1383  | 
apply (fast dest: approx_subset_monad)  | 
|
1384  | 
done  | 
|
| 27468 | 1385  | 
|
| 64435 | 1386  | 
lemma mem_monad_approx: "u \<in> monad x \<Longrightarrow> x \<approx> u"  | 
1387  | 
by (simp add: monad_def)  | 
|
1388  | 
||
1389  | 
lemma approx_mem_monad: "x \<approx> u \<Longrightarrow> u \<in> monad x"  | 
|
1390  | 
by (simp add: monad_def)  | 
|
| 27468 | 1391  | 
|
| 64435 | 1392  | 
lemma approx_mem_monad2: "x \<approx> u \<Longrightarrow> x \<in> monad u"  | 
1393  | 
apply (simp add: monad_def)  | 
|
1394  | 
apply (blast intro!: approx_sym)  | 
|
1395  | 
done  | 
|
| 27468 | 1396  | 
|
| 64435 | 1397  | 
lemma approx_mem_monad_zero: "x \<approx> y \<Longrightarrow> x \<in> monad 0 \<Longrightarrow> y \<in> monad 0"  | 
1398  | 
apply (drule mem_monad_approx)  | 
|
1399  | 
apply (fast intro: approx_mem_monad approx_trans)  | 
|
1400  | 
done  | 
|
| 27468 | 1401  | 
|
| 64435 | 1402  | 
lemma Infinitesimal_approx_hrabs: "x \<approx> y \<Longrightarrow> x \<in> Infinitesimal \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>y\<bar>"  | 
1403  | 
for x y :: hypreal  | 
|
1404  | 
apply (drule Infinitesimal_monad_zero_iff [THEN iffD1])  | 
|
1405  | 
apply (blast intro: approx_mem_monad_zero monad_zero_hrabs_iff [THEN iffD1]  | 
|
1406  | 
mem_monad_approx approx_trans3)  | 
|
1407  | 
done  | 
|
| 27468 | 1408  | 
|
| 64435 | 1409  | 
lemma less_Infinitesimal_less: "0 < x \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> e \<in> Infinitesimal \<Longrightarrow> e < x"  | 
1410  | 
for x :: hypreal  | 
|
1411  | 
apply (rule ccontr)  | 
|
1412  | 
apply (auto intro: Infinitesimal_zero [THEN [2] Infinitesimal_interval]  | 
|
1413  | 
dest!: order_le_imp_less_or_eq simp add: linorder_not_less)  | 
|
1414  | 
done  | 
|
| 27468 | 1415  | 
|
| 64435 | 1416  | 
lemma Ball_mem_monad_gt_zero: "0 < x \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> u \<in> monad x \<Longrightarrow> 0 < u"  | 
1417  | 
for u x :: hypreal  | 
|
1418  | 
apply (drule mem_monad_approx [THEN approx_sym])  | 
|
1419  | 
apply (erule bex_Infinitesimal_iff2 [THEN iffD2, THEN bexE])  | 
|
1420  | 
apply (drule_tac e = "-xa" in less_Infinitesimal_less, auto)  | 
|
1421  | 
done  | 
|
| 27468 | 1422  | 
|
| 64435 | 1423  | 
lemma Ball_mem_monad_less_zero: "x < 0 \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> u \<in> monad x \<Longrightarrow> u < 0"  | 
1424  | 
for u x :: hypreal  | 
|
1425  | 
apply (drule mem_monad_approx [THEN approx_sym])  | 
|
1426  | 
apply (erule bex_Infinitesimal_iff [THEN iffD2, THEN bexE])  | 
|
1427  | 
apply (cut_tac x = "-x" and e = xa in less_Infinitesimal_less, auto)  | 
|
1428  | 
done  | 
|
| 27468 | 1429  | 
|
| 64435 | 1430  | 
lemma lemma_approx_gt_zero: "0 < x \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> x \<approx> y \<Longrightarrow> 0 < y"  | 
1431  | 
for x y :: hypreal  | 
|
1432  | 
by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad)  | 
|
| 27468 | 1433  | 
|
| 64435 | 1434  | 
lemma lemma_approx_less_zero: "x < 0 \<Longrightarrow> x \<notin> Infinitesimal \<Longrightarrow> x \<approx> y \<Longrightarrow> y < 0"  | 
1435  | 
for x y :: hypreal  | 
|
1436  | 
by (blast dest: Ball_mem_monad_less_zero approx_subset_monad)  | 
|
| 27468 | 1437  | 
|
| 64435 | 1438  | 
lemma approx_hrabs: "x \<approx> y \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>y\<bar>"  | 
1439  | 
for x y :: hypreal  | 
|
1440  | 
by (drule approx_hnorm) simp  | 
|
| 27468 | 1441  | 
|
| 64435 | 1442  | 
lemma approx_hrabs_zero_cancel: "\<bar>x\<bar> \<approx> 0 \<Longrightarrow> x \<approx> 0"  | 
1443  | 
for x :: hypreal  | 
|
1444  | 
using hrabs_disj [of x] by (auto dest: approx_minus)  | 
|
| 27468 | 1445  | 
|
| 64435 | 1446  | 
lemma approx_hrabs_add_Infinitesimal: "e \<in> Infinitesimal \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>x + e\<bar>"  | 
1447  | 
for e x :: hypreal  | 
|
1448  | 
by (fast intro: approx_hrabs Infinitesimal_add_approx_self)  | 
|
| 27468 | 1449  | 
|
| 64435 | 1450  | 
lemma approx_hrabs_add_minus_Infinitesimal: "e \<in> Infinitesimal ==> \<bar>x\<bar> \<approx> \<bar>x + -e\<bar>"  | 
1451  | 
for e x :: hypreal  | 
|
1452  | 
by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self)  | 
|
| 27468 | 1453  | 
|
1454  | 
lemma hrabs_add_Infinitesimal_cancel:  | 
|
| 64435 | 1455  | 
"e \<in> Infinitesimal \<Longrightarrow> e' \<in> Infinitesimal \<Longrightarrow> \<bar>x + e\<bar> = \<bar>y + e'\<bar> \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>y\<bar>"  | 
1456  | 
for e e' x y :: hypreal  | 
|
1457  | 
apply (drule_tac x = x in approx_hrabs_add_Infinitesimal)  | 
|
1458  | 
apply (drule_tac x = y in approx_hrabs_add_Infinitesimal)  | 
|
1459  | 
apply (auto intro: approx_trans2)  | 
|
1460  | 
done  | 
|
| 27468 | 1461  | 
|
1462  | 
lemma hrabs_add_minus_Infinitesimal_cancel:  | 
|
| 64435 | 1463  | 
"e \<in> Infinitesimal \<Longrightarrow> e' \<in> Infinitesimal \<Longrightarrow> \<bar>x + -e\<bar> = \<bar>y + -e'\<bar> \<Longrightarrow> \<bar>x\<bar> \<approx> \<bar>y\<bar>"  | 
1464  | 
for e e' x y :: hypreal  | 
|
1465  | 
apply (drule_tac x = x in approx_hrabs_add_minus_Infinitesimal)  | 
|
1466  | 
apply (drule_tac x = y in approx_hrabs_add_minus_Infinitesimal)  | 
|
1467  | 
apply (auto intro: approx_trans2)  | 
|
1468  | 
done  | 
|
1469  | 
||
| 27468 | 1470  | 
|
| 61975 | 1471  | 
subsection \<open>More @{term HFinite} and @{term Infinitesimal} Theorems\<close>
 | 
| 27468 | 1472  | 
|
| 64435 | 1473  | 
text \<open>  | 
1474  | 
Interesting slightly counterintuitive theorem: necessary  | 
|
1475  | 
for proving that an open interval is an NS open set.  | 
|
1476  | 
\<close>  | 
|
| 27468 | 1477  | 
lemma Infinitesimal_add_hypreal_of_real_less:  | 
| 64435 | 1478  | 
"x < y \<Longrightarrow> u \<in> Infinitesimal \<Longrightarrow> hypreal_of_real x + u < hypreal_of_real y"  | 
1479  | 
apply (simp add: Infinitesimal_def)  | 
|
1480  | 
apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec, simp)  | 
|
1481  | 
apply (simp add: abs_less_iff)  | 
|
1482  | 
done  | 
|
| 27468 | 1483  | 
|
1484  | 
lemma Infinitesimal_add_hrabs_hypreal_of_real_less:  | 
|
| 64435 | 1485  | 
"x \<in> Infinitesimal \<Longrightarrow> \<bar>hypreal_of_real r\<bar> < hypreal_of_real y \<Longrightarrow>  | 
1486  | 
\<bar>hypreal_of_real r + x\<bar> < hypreal_of_real y"  | 
|
1487  | 
apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal)  | 
|
1488  | 
apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]])  | 
|
1489  | 
apply (auto intro!: Infinitesimal_add_hypreal_of_real_less  | 
|
1490  | 
simp del: star_of_abs simp add: star_of_abs [symmetric])  | 
|
1491  | 
done  | 
|
| 27468 | 1492  | 
|
1493  | 
lemma Infinitesimal_add_hrabs_hypreal_of_real_less2:  | 
|
| 64435 | 1494  | 
"x \<in> Infinitesimal \<Longrightarrow> \<bar>hypreal_of_real r\<bar> < hypreal_of_real y \<Longrightarrow>  | 
1495  | 
\<bar>x + hypreal_of_real r\<bar> < hypreal_of_real y"  | 
|
1496  | 
apply (rule add.commute [THEN subst])  | 
|
1497  | 
apply (erule Infinitesimal_add_hrabs_hypreal_of_real_less, assumption)  | 
|
1498  | 
done  | 
|
| 27468 | 1499  | 
|
1500  | 
lemma hypreal_of_real_le_add_Infininitesimal_cancel:  | 
|
| 64435 | 1501  | 
"u \<in> Infinitesimal \<Longrightarrow> v \<in> Infinitesimal \<Longrightarrow>  | 
1502  | 
hypreal_of_real x + u \<le> hypreal_of_real y + v \<Longrightarrow>  | 
|
1503  | 
hypreal_of_real x \<le> hypreal_of_real y"  | 
|
1504  | 
apply (simp add: linorder_not_less [symmetric], auto)  | 
|
1505  | 
apply (drule_tac u = "v-u" in Infinitesimal_add_hypreal_of_real_less)  | 
|
1506  | 
apply (auto simp add: Infinitesimal_diff)  | 
|
1507  | 
done  | 
|
| 27468 | 1508  | 
|
1509  | 
lemma hypreal_of_real_le_add_Infininitesimal_cancel2:  | 
|
| 64435 | 1510  | 
"u \<in> Infinitesimal \<Longrightarrow> v \<in> Infinitesimal \<Longrightarrow>  | 
1511  | 
hypreal_of_real x + u \<le> hypreal_of_real y + v \<Longrightarrow> x \<le> y"  | 
|
1512  | 
by (blast intro: star_of_le [THEN iffD1] intro!: hypreal_of_real_le_add_Infininitesimal_cancel)  | 
|
| 27468 | 1513  | 
|
1514  | 
lemma hypreal_of_real_less_Infinitesimal_le_zero:  | 
|
| 64435 | 1515  | 
"hypreal_of_real x < e \<Longrightarrow> e \<in> Infinitesimal \<Longrightarrow> hypreal_of_real x \<le> 0"  | 
1516  | 
apply (rule linorder_not_less [THEN iffD1], safe)  | 
|
1517  | 
apply (drule Infinitesimal_interval)  | 
|
1518  | 
apply (drule_tac [4] SReal_hypreal_of_real [THEN SReal_Infinitesimal_zero], auto)  | 
|
1519  | 
done  | 
|
| 27468 | 1520  | 
|
1521  | 
(*used once, in Lim/NSDERIV_inverse*)  | 
|
| 64435 | 1522  | 
lemma Infinitesimal_add_not_zero: "h \<in> Infinitesimal \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> star_of x + h \<noteq> 0"  | 
1523  | 
apply auto  | 
|
1524  | 
apply (subgoal_tac "h = - star_of x")  | 
|
1525  | 
apply (auto intro: minus_unique [symmetric])  | 
|
1526  | 
done  | 
|
| 27468 | 1527  | 
|
| 64435 | 1528  | 
lemma Infinitesimal_square_cancel [simp]: "x * x + y * y \<in> Infinitesimal \<Longrightarrow> x * x \<in> Infinitesimal"  | 
1529  | 
for x y :: hypreal  | 
|
1530  | 
apply (rule Infinitesimal_interval2)  | 
|
1531  | 
apply (rule_tac [3] zero_le_square, assumption)  | 
|
1532  | 
apply auto  | 
|
1533  | 
done  | 
|
| 27468 | 1534  | 
|
| 64435 | 1535  | 
lemma HFinite_square_cancel [simp]: "x * x + y * y \<in> HFinite \<Longrightarrow> x * x \<in> HFinite"  | 
1536  | 
for x y :: hypreal  | 
|
1537  | 
apply (rule HFinite_bounded, assumption)  | 
|
1538  | 
apply auto  | 
|
1539  | 
done  | 
|
| 27468 | 1540  | 
|
| 64435 | 1541  | 
lemma Infinitesimal_square_cancel2 [simp]: "x * x + y * y \<in> Infinitesimal \<Longrightarrow> y * y \<in> Infinitesimal"  | 
1542  | 
for x y :: hypreal  | 
|
1543  | 
apply (rule Infinitesimal_square_cancel)  | 
|
1544  | 
apply (rule add.commute [THEN subst])  | 
|
1545  | 
apply simp  | 
|
1546  | 
done  | 
|
| 27468 | 1547  | 
|
| 64435 | 1548  | 
lemma HFinite_square_cancel2 [simp]: "x * x + y * y \<in> HFinite \<Longrightarrow> y * y \<in> HFinite"  | 
1549  | 
for x y :: hypreal  | 
|
1550  | 
apply (rule HFinite_square_cancel)  | 
|
1551  | 
apply (rule add.commute [THEN subst])  | 
|
1552  | 
apply simp  | 
|
1553  | 
done  | 
|
| 27468 | 1554  | 
|
1555  | 
lemma Infinitesimal_sum_square_cancel [simp]:  | 
|
| 64435 | 1556  | 
"x * x + y * y + z * z \<in> Infinitesimal \<Longrightarrow> x * x \<in> Infinitesimal"  | 
1557  | 
for x y z :: hypreal  | 
|
1558  | 
apply (rule Infinitesimal_interval2, assumption)  | 
|
1559  | 
apply (rule_tac [2] zero_le_square, simp)  | 
|
1560  | 
apply (insert zero_le_square [of y])  | 
|
1561  | 
apply (insert zero_le_square [of z], simp del:zero_le_square)  | 
|
1562  | 
done  | 
|
| 27468 | 1563  | 
|
| 64435 | 1564  | 
lemma HFinite_sum_square_cancel [simp]: "x * x + y * y + z * z \<in> HFinite \<Longrightarrow> x * x \<in> HFinite"  | 
1565  | 
for x y z :: hypreal  | 
|
1566  | 
apply (rule HFinite_bounded, assumption)  | 
|
1567  | 
apply (rule_tac [2] zero_le_square)  | 
|
1568  | 
apply (insert zero_le_square [of y])  | 
|
1569  | 
apply (insert zero_le_square [of z], simp del:zero_le_square)  | 
|
1570  | 
done  | 
|
| 27468 | 1571  | 
|
1572  | 
lemma Infinitesimal_sum_square_cancel2 [simp]:  | 
|
| 64435 | 1573  | 
"y * y + x * x + z * z \<in> Infinitesimal \<Longrightarrow> x * x \<in> Infinitesimal"  | 
1574  | 
for x y z :: hypreal  | 
|
1575  | 
apply (rule Infinitesimal_sum_square_cancel)  | 
|
1576  | 
apply (simp add: ac_simps)  | 
|
1577  | 
done  | 
|
| 27468 | 1578  | 
|
| 64435 | 1579  | 
lemma HFinite_sum_square_cancel2 [simp]: "y * y + x * x + z * z \<in> HFinite \<Longrightarrow> x * x \<in> HFinite"  | 
1580  | 
for x y z :: hypreal  | 
|
1581  | 
apply (rule HFinite_sum_square_cancel)  | 
|
1582  | 
apply (simp add: ac_simps)  | 
|
1583  | 
done  | 
|
| 27468 | 1584  | 
|
1585  | 
lemma Infinitesimal_sum_square_cancel3 [simp]:  | 
|
| 64435 | 1586  | 
"z * z + y * y + x * x \<in> Infinitesimal \<Longrightarrow> x * x \<in> Infinitesimal"  | 
1587  | 
for x y z :: hypreal  | 
|
1588  | 
apply (rule Infinitesimal_sum_square_cancel)  | 
|
1589  | 
apply (simp add: ac_simps)  | 
|
1590  | 
done  | 
|
| 27468 | 1591  | 
|
| 64435 | 1592  | 
lemma HFinite_sum_square_cancel3 [simp]: "z * z + y * y + x * x \<in> HFinite \<Longrightarrow> x * x \<in> HFinite"  | 
1593  | 
for x y z :: hypreal  | 
|
1594  | 
apply (rule HFinite_sum_square_cancel)  | 
|
1595  | 
apply (simp add: ac_simps)  | 
|
1596  | 
done  | 
|
| 27468 | 1597  | 
|
| 64435 | 1598  | 
lemma monad_hrabs_less: "y \<in> monad x \<Longrightarrow> 0 < hypreal_of_real e \<Longrightarrow> \<bar>y - x\<bar> < hypreal_of_real e"  | 
1599  | 
apply (drule mem_monad_approx [THEN approx_sym])  | 
|
1600  | 
apply (drule bex_Infinitesimal_iff [THEN iffD2])  | 
|
1601  | 
apply (auto dest!: InfinitesimalD)  | 
|
1602  | 
done  | 
|
| 27468 | 1603  | 
|
| 64435 | 1604  | 
lemma mem_monad_SReal_HFinite: "x \<in> monad (hypreal_of_real a) \<Longrightarrow> x \<in> HFinite"  | 
1605  | 
apply (drule mem_monad_approx [THEN approx_sym])  | 
|
1606  | 
apply (drule bex_Infinitesimal_iff2 [THEN iffD2])  | 
|
1607  | 
apply (safe dest!: Infinitesimal_subset_HFinite [THEN subsetD])  | 
|
1608  | 
apply (erule SReal_hypreal_of_real [THEN SReal_subset_HFinite [THEN subsetD], THEN HFinite_add])  | 
|
1609  | 
done  | 
|
| 27468 | 1610  | 
|
1611  | 
||
| 64435 | 1612  | 
subsection \<open>Theorems about Standard Part\<close>  | 
| 27468 | 1613  | 
|
| 64435 | 1614  | 
lemma st_approx_self: "x \<in> HFinite \<Longrightarrow> st x \<approx> x"  | 
1615  | 
apply (simp add: st_def)  | 
|
1616  | 
apply (frule st_part_Ex, safe)  | 
|
1617  | 
apply (rule someI2)  | 
|
1618  | 
apply (auto intro: approx_sym)  | 
|
1619  | 
done  | 
|
| 27468 | 1620  | 
|
| 64435 | 1621  | 
lemma st_SReal: "x \<in> HFinite \<Longrightarrow> st x \<in> \<real>"  | 
1622  | 
apply (simp add: st_def)  | 
|
1623  | 
apply (frule st_part_Ex, safe)  | 
|
1624  | 
apply (rule someI2)  | 
|
1625  | 
apply (auto intro: approx_sym)  | 
|
1626  | 
done  | 
|
| 27468 | 1627  | 
|
| 64435 | 1628  | 
lemma st_HFinite: "x \<in> HFinite \<Longrightarrow> st x \<in> HFinite"  | 
1629  | 
by (erule st_SReal [THEN SReal_subset_HFinite [THEN subsetD]])  | 
|
| 27468 | 1630  | 
|
| 64435 | 1631  | 
lemma st_unique: "r \<in> \<real> \<Longrightarrow> r \<approx> x \<Longrightarrow> st x = r"  | 
1632  | 
apply (frule SReal_subset_HFinite [THEN subsetD])  | 
|
1633  | 
apply (drule (1) approx_HFinite)  | 
|
1634  | 
apply (unfold st_def)  | 
|
1635  | 
apply (rule some_equality)  | 
|
1636  | 
apply (auto intro: approx_unique_real)  | 
|
1637  | 
done  | 
|
| 27468 | 1638  | 
|
| 64435 | 1639  | 
lemma st_SReal_eq: "x \<in> \<real> \<Longrightarrow> st x = x"  | 
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61609 
diff
changeset
 | 
1640  | 
by (metis approx_refl st_unique)  | 
| 27468 | 1641  | 
|
1642  | 
lemma st_hypreal_of_real [simp]: "st (hypreal_of_real x) = hypreal_of_real x"  | 
|
| 64435 | 1643  | 
by (rule SReal_hypreal_of_real [THEN st_SReal_eq])  | 
| 27468 | 1644  | 
|
| 64435 | 1645  | 
lemma st_eq_approx: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> st x = st y \<Longrightarrow> x \<approx> y"  | 
1646  | 
by (auto dest!: st_approx_self elim!: approx_trans3)  | 
|
| 27468 | 1647  | 
|
| 
56217
 
dc429a5b13c4
Some rationalisation of basic lemmas
 
paulson <lp15@cam.ac.uk> 
parents: 
54489 
diff
changeset
 | 
1648  | 
lemma approx_st_eq:  | 
| 61982 | 1649  | 
assumes x: "x \<in> HFinite" and y: "y \<in> HFinite" and xy: "x \<approx> y"  | 
| 27468 | 1650  | 
shows "st x = st y"  | 
1651  | 
proof -  | 
|
| 61982 | 1652  | 
have "st x \<approx> x" "st y \<approx> y" "st x \<in> \<real>" "st y \<in> \<real>"  | 
| 41541 | 1653  | 
by (simp_all add: st_approx_self st_SReal x y)  | 
1654  | 
with xy show ?thesis  | 
|
| 27468 | 1655  | 
by (fast elim: approx_trans approx_trans2 SReal_approx_iff [THEN iffD1])  | 
1656  | 
qed  | 
|
1657  | 
||
| 64435 | 1658  | 
lemma st_eq_approx_iff: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x \<approx> y \<longleftrightarrow> st x = st y"  | 
1659  | 
by (blast intro: approx_st_eq st_eq_approx)  | 
|
| 27468 | 1660  | 
|
| 64435 | 1661  | 
lemma st_Infinitesimal_add_SReal: "x \<in> \<real> \<Longrightarrow> e \<in> Infinitesimal \<Longrightarrow> st (x + e) = x"  | 
1662  | 
apply (erule st_unique)  | 
|
1663  | 
apply (erule Infinitesimal_add_approx_self)  | 
|
1664  | 
done  | 
|
| 27468 | 1665  | 
|
| 64435 | 1666  | 
lemma st_Infinitesimal_add_SReal2: "x \<in> \<real> \<Longrightarrow> e \<in> Infinitesimal \<Longrightarrow> st (e + x) = x"  | 
1667  | 
apply (erule st_unique)  | 
|
1668  | 
apply (erule Infinitesimal_add_approx_self2)  | 
|
1669  | 
done  | 
|
| 27468 | 1670  | 
|
| 64435 | 1671  | 
lemma HFinite_st_Infinitesimal_add: "x \<in> HFinite \<Longrightarrow> \<exists>e \<in> Infinitesimal. x = st(x) + e"  | 
1672  | 
by (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2])  | 
|
| 27468 | 1673  | 
|
| 64435 | 1674  | 
lemma st_add: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> st (x + y) = st x + st y"  | 
1675  | 
by (simp add: st_unique st_SReal st_approx_self approx_add)  | 
|
| 27468 | 1676  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
45541 
diff
changeset
 | 
1677  | 
lemma st_numeral [simp]: "st (numeral w) = numeral w"  | 
| 64435 | 1678  | 
by (rule Reals_numeral [THEN st_SReal_eq])  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
45541 
diff
changeset
 | 
1679  | 
|
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
1680  | 
lemma st_neg_numeral [simp]: "st (- numeral w) = - numeral w"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
1681  | 
proof -  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
1682  | 
from Reals_numeral have "numeral w \<in> \<real>" .  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
1683  | 
then have "- numeral w \<in> \<real>" by simp  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
1684  | 
with st_SReal_eq show ?thesis .  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
1685  | 
qed  | 
| 27468 | 1686  | 
|
| 45540 | 1687  | 
lemma st_0 [simp]: "st 0 = 0"  | 
| 64435 | 1688  | 
by (simp add: st_SReal_eq)  | 
| 45540 | 1689  | 
|
1690  | 
lemma st_1 [simp]: "st 1 = 1"  | 
|
| 64435 | 1691  | 
by (simp add: st_SReal_eq)  | 
| 27468 | 1692  | 
|
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
1693  | 
lemma st_neg_1 [simp]: "st (- 1) = - 1"  | 
| 64435 | 1694  | 
by (simp add: st_SReal_eq)  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
1695  | 
|
| 27468 | 1696  | 
lemma st_minus: "x \<in> HFinite \<Longrightarrow> st (- x) = - st x"  | 
| 64435 | 1697  | 
by (simp add: st_unique st_SReal st_approx_self approx_minus)  | 
| 27468 | 1698  | 
|
1699  | 
lemma st_diff: "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> st (x - y) = st x - st y"  | 
|
| 64435 | 1700  | 
by (simp add: st_unique st_SReal st_approx_self approx_diff)  | 
| 27468 | 1701  | 
|
1702  | 
lemma st_mult: "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> st (x * y) = st x * st y"  | 
|
| 64435 | 1703  | 
by (simp add: st_unique st_SReal st_approx_self approx_mult_HFinite)  | 
| 27468 | 1704  | 
|
| 64435 | 1705  | 
lemma st_Infinitesimal: "x \<in> Infinitesimal \<Longrightarrow> st x = 0"  | 
1706  | 
by (simp add: st_unique mem_infmal_iff)  | 
|
| 27468 | 1707  | 
|
| 64435 | 1708  | 
lemma st_not_Infinitesimal: "st(x) \<noteq> 0 \<Longrightarrow> x \<notin> Infinitesimal"  | 
| 27468 | 1709  | 
by (fast intro: st_Infinitesimal)  | 
1710  | 
||
| 64435 | 1711  | 
lemma st_inverse: "x \<in> HFinite \<Longrightarrow> st x \<noteq> 0 \<Longrightarrow> st (inverse x) = inverse (st x)"  | 
1712  | 
apply (rule_tac c1 = "st x" in mult_left_cancel [THEN iffD1])  | 
|
1713  | 
apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse)  | 
|
1714  | 
apply (subst right_inverse, auto)  | 
|
1715  | 
done  | 
|
| 27468 | 1716  | 
|
| 64435 | 1717  | 
lemma st_divide [simp]: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> st y \<noteq> 0 \<Longrightarrow> st (x / y) = st x / st y"  | 
1718  | 
by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse)  | 
|
| 27468 | 1719  | 
|
| 64435 | 1720  | 
lemma st_idempotent [simp]: "x \<in> HFinite \<Longrightarrow> st (st x) = st x"  | 
1721  | 
by (blast intro: st_HFinite st_approx_self approx_st_eq)  | 
|
| 27468 | 1722  | 
|
1723  | 
lemma Infinitesimal_add_st_less:  | 
|
| 64435 | 1724  | 
"x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> u \<in> Infinitesimal \<Longrightarrow> st x < st y \<Longrightarrow> st x + u < st y"  | 
1725  | 
apply (drule st_SReal)+  | 
|
1726  | 
apply (auto intro!: Infinitesimal_add_hypreal_of_real_less simp add: SReal_iff)  | 
|
1727  | 
done  | 
|
| 27468 | 1728  | 
|
1729  | 
lemma Infinitesimal_add_st_le_cancel:  | 
|
| 64435 | 1730  | 
"x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> u \<in> Infinitesimal \<Longrightarrow>  | 
1731  | 
st x \<le> st y + u \<Longrightarrow> st x \<le> st y"  | 
|
1732  | 
apply (simp add: linorder_not_less [symmetric])  | 
|
1733  | 
apply (auto dest: Infinitesimal_add_st_less)  | 
|
1734  | 
done  | 
|
| 27468 | 1735  | 
|
| 64435 | 1736  | 
lemma st_le: "x \<in> HFinite \<Longrightarrow> y \<in> HFinite \<Longrightarrow> x \<le> y \<Longrightarrow> st x \<le> st y"  | 
1737  | 
by (metis approx_le_bound approx_sym linear st_SReal st_approx_self st_part_Ex1)  | 
|
| 27468 | 1738  | 
|
| 64435 | 1739  | 
lemma st_zero_le: "0 \<le> x \<Longrightarrow> x \<in> HFinite \<Longrightarrow> 0 \<le> st x"  | 
1740  | 
apply (subst st_0 [symmetric])  | 
|
1741  | 
apply (rule st_le, auto)  | 
|
1742  | 
done  | 
|
| 27468 | 1743  | 
|
| 64435 | 1744  | 
lemma st_zero_ge: "x \<le> 0 \<Longrightarrow> x \<in> HFinite \<Longrightarrow> st x \<le> 0"  | 
1745  | 
apply (subst st_0 [symmetric])  | 
|
1746  | 
apply (rule st_le, auto)  | 
|
1747  | 
done  | 
|
| 27468 | 1748  | 
|
| 64435 | 1749  | 
lemma st_hrabs: "x \<in> HFinite \<Longrightarrow> \<bar>st x\<bar> = st \<bar>x\<bar>"  | 
1750  | 
apply (simp add: linorder_not_le st_zero_le abs_if st_minus linorder_not_less)  | 
|
1751  | 
apply (auto dest!: st_zero_ge [OF order_less_imp_le])  | 
|
1752  | 
done  | 
|
| 27468 | 1753  | 
|
1754  | 
||
| 61975 | 1755  | 
subsection \<open>Alternative Definitions using Free Ultrafilter\<close>  | 
| 27468 | 1756  | 
|
| 61975 | 1757  | 
subsubsection \<open>@{term HFinite}\<close>
 | 
| 27468 | 1758  | 
|
1759  | 
lemma HFinite_FreeUltrafilterNat:  | 
|
| 64438 | 1760  | 
"star_n X \<in> HFinite \<Longrightarrow> \<exists>u. eventually (\<lambda>n. norm (X n) < u) \<U>"  | 
| 64435 | 1761  | 
apply (auto simp add: HFinite_def SReal_def)  | 
1762  | 
apply (rule_tac x=r in exI)  | 
|
1763  | 
apply (simp add: hnorm_def star_of_def starfun_star_n)  | 
|
1764  | 
apply (simp add: star_less_def starP2_star_n)  | 
|
1765  | 
done  | 
|
| 27468 | 1766  | 
|
1767  | 
lemma FreeUltrafilterNat_HFinite:  | 
|
| 64438 | 1768  | 
"\<exists>u. eventually (\<lambda>n. norm (X n) < u) \<U> \<Longrightarrow> star_n X \<in> HFinite"  | 
| 64435 | 1769  | 
apply (auto simp add: HFinite_def mem_Rep_star_iff)  | 
1770  | 
apply (rule_tac x="star_of u" in bexI)  | 
|
1771  | 
apply (simp add: hnorm_def starfun_star_n star_of_def)  | 
|
1772  | 
apply (simp add: star_less_def starP2_star_n)  | 
|
1773  | 
apply (simp add: SReal_def)  | 
|
1774  | 
done  | 
|
| 27468 | 1775  | 
|
1776  | 
lemma HFinite_FreeUltrafilterNat_iff:  | 
|
| 64438 | 1777  | 
"star_n X \<in> HFinite \<longleftrightarrow> (\<exists>u. eventually (\<lambda>n. norm (X n) < u) \<U>)"  | 
| 64435 | 1778  | 
by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite)  | 
1779  | 
||
| 27468 | 1780  | 
|
| 61975 | 1781  | 
subsubsection \<open>@{term HInfinite}\<close>
 | 
| 27468 | 1782  | 
|
| 56225 | 1783  | 
lemma lemma_Compl_eq: "- {n. u < norm (f n)} = {n. norm (f n) \<le> u}"
 | 
| 64435 | 1784  | 
by auto  | 
| 27468 | 1785  | 
|
| 56225 | 1786  | 
lemma lemma_Compl_eq2: "- {n. norm (f n) < u} = {n. u \<le> norm (f n)}"
 | 
| 64435 | 1787  | 
by auto  | 
| 27468 | 1788  | 
|
| 64435 | 1789  | 
lemma lemma_Int_eq1: "{n. norm (f n) \<le> u} Int {n. u \<le> norm (f n)} = {n. norm(f n) = u}"
 | 
1790  | 
by auto  | 
|
| 27468 | 1791  | 
|
| 64435 | 1792  | 
lemma lemma_FreeUltrafilterNat_one: "{n. norm (f n) = u} \<le> {n. norm (f n) < u + (1::real)}"
 | 
1793  | 
by auto  | 
|
| 27468 | 1794  | 
|
| 64435 | 1795  | 
text \<open>Exclude this type of sets from free ultrafilter for Infinite numbers!\<close>  | 
| 27468 | 1796  | 
lemma FreeUltrafilterNat_const_Finite:  | 
| 64438 | 1797  | 
"eventually (\<lambda>n. norm (X n) = u) \<U> \<Longrightarrow> star_n X \<in> HFinite"  | 
| 64435 | 1798  | 
apply (rule FreeUltrafilterNat_HFinite)  | 
1799  | 
apply (rule_tac x = "u + 1" in exI)  | 
|
1800  | 
apply (auto elim: eventually_mono)  | 
|
1801  | 
done  | 
|
| 27468 | 1802  | 
|
1803  | 
lemma HInfinite_FreeUltrafilterNat:  | 
|
| 64438 | 1804  | 
"star_n X \<in> HInfinite \<Longrightarrow> \<forall>u. eventually (\<lambda>n. u < norm (X n)) \<U>"  | 
| 64435 | 1805  | 
apply (drule HInfinite_HFinite_iff [THEN iffD1])  | 
1806  | 
apply (simp add: HFinite_FreeUltrafilterNat_iff)  | 
|
1807  | 
apply (rule allI, drule_tac x="u + 1" in spec)  | 
|
1808  | 
apply (simp add: FreeUltrafilterNat.eventually_not_iff[symmetric])  | 
|
1809  | 
apply (auto elim: eventually_mono)  | 
|
1810  | 
done  | 
|
| 27468 | 1811  | 
|
| 64435 | 1812  | 
lemma lemma_Int_HI: "{n. norm (Xa n) < u} \<inter> {n. X n = Xa n} \<subseteq> {n. norm (X n) < u}"
 | 
1813  | 
for u :: real  | 
|
1814  | 
by auto  | 
|
| 27468 | 1815  | 
|
| 64435 | 1816  | 
lemma lemma_Int_HIa: "{n. u < norm (X n)} \<inter> {n. norm (X n) < u} = {}"
 | 
1817  | 
by (auto intro: order_less_asym)  | 
|
| 27468 | 1818  | 
|
1819  | 
lemma FreeUltrafilterNat_HInfinite:  | 
|
| 64438 | 1820  | 
"\<forall>u. eventually (\<lambda>n. u < norm (X n)) \<U> \<Longrightarrow> star_n X \<in> HInfinite"  | 
| 64435 | 1821  | 
apply (rule HInfinite_HFinite_iff [THEN iffD2])  | 
1822  | 
apply (safe, drule HFinite_FreeUltrafilterNat, safe)  | 
|
1823  | 
apply (drule_tac x = u in spec)  | 
|
| 60041 | 1824  | 
proof -  | 
| 64435 | 1825  | 
fix u  | 
1826  | 
assume "\<forall>\<^sub>Fn in \<U>. norm (X n) < u" "\<forall>\<^sub>Fn in \<U>. u < norm (X n)"  | 
|
| 60041 | 1827  | 
then have "\<forall>\<^sub>F x in \<U>. False"  | 
1828  | 
by eventually_elim auto  | 
|
1829  | 
then show False  | 
|
1830  | 
by (simp add: eventually_False FreeUltrafilterNat.proper)  | 
|
1831  | 
qed  | 
|
| 27468 | 1832  | 
|
1833  | 
lemma HInfinite_FreeUltrafilterNat_iff:  | 
|
| 64438 | 1834  | 
"star_n X \<in> HInfinite \<longleftrightarrow> (\<forall>u. eventually (\<lambda>n. u < norm (X n)) \<U>)"  | 
| 64435 | 1835  | 
by (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite)  | 
1836  | 
||
| 27468 | 1837  | 
|
| 61975 | 1838  | 
subsubsection \<open>@{term Infinitesimal}\<close>
 | 
| 27468 | 1839  | 
|
| 64435 | 1840  | 
lemma ball_SReal_eq: "(\<forall>x::hypreal \<in> Reals. P x) \<longleftrightarrow> (\<forall>x::real. P (star_of x))"  | 
1841  | 
by (auto simp: SReal_def)  | 
|
| 27468 | 1842  | 
|
1843  | 
lemma Infinitesimal_FreeUltrafilterNat:  | 
|
| 64435 | 1844  | 
"star_n X \<in> Infinitesimal \<Longrightarrow> \<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U>"  | 
1845  | 
apply (simp add: Infinitesimal_def ball_SReal_eq)  | 
|
1846  | 
apply (simp add: hnorm_def starfun_star_n star_of_def)  | 
|
1847  | 
apply (simp add: star_less_def starP2_star_n)  | 
|
1848  | 
done  | 
|
| 27468 | 1849  | 
|
1850  | 
lemma FreeUltrafilterNat_Infinitesimal:  | 
|
| 64435 | 1851  | 
"\<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U> \<Longrightarrow> star_n X \<in> Infinitesimal"  | 
1852  | 
apply (simp add: Infinitesimal_def ball_SReal_eq)  | 
|
1853  | 
apply (simp add: hnorm_def starfun_star_n star_of_def)  | 
|
1854  | 
apply (simp add: star_less_def starP2_star_n)  | 
|
1855  | 
done  | 
|
| 27468 | 1856  | 
|
1857  | 
lemma Infinitesimal_FreeUltrafilterNat_iff:  | 
|
| 64435 | 1858  | 
"(star_n X \<in> Infinitesimal) = (\<forall>u>0. eventually (\<lambda>n. norm (X n) < u) \<U>)"  | 
1859  | 
by (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal)  | 
|
1860  | 
||
| 27468 | 1861  | 
|
| 64435 | 1862  | 
text \<open>Infinitesimals as smaller than \<open>1/n\<close> for all \<open>n::nat (> 0)\<close>.\<close>  | 
| 27468 | 1863  | 
|
| 64435 | 1864  | 
lemma lemma_Infinitesimal: "(\<forall>r. 0 < r \<longrightarrow> x < r) \<longleftrightarrow> (\<forall>n. x < inverse (real (Suc n)))"  | 
1865  | 
apply (auto simp del: of_nat_Suc)  | 
|
1866  | 
apply (blast dest!: reals_Archimedean intro: order_less_trans)  | 
|
1867  | 
done  | 
|
| 27468 | 1868  | 
|
1869  | 
lemma lemma_Infinitesimal2:  | 
|
| 64435 | 1870  | 
"(\<forall>r \<in> Reals. 0 < r \<longrightarrow> x < r) \<longleftrightarrow> (\<forall>n. x < inverse(hypreal_of_nat (Suc n)))"  | 
1871  | 
apply safe  | 
|
1872  | 
apply (drule_tac x = "inverse (hypreal_of_real (real (Suc n))) " in bspec)  | 
|
1873  | 
apply simp_all  | 
|
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61609 
diff
changeset
 | 
1874  | 
using less_imp_of_nat_less apply fastforce  | 
| 64435 | 1875  | 
apply (auto dest!: reals_Archimedean simp add: SReal_iff simp del: of_nat_Suc)  | 
1876  | 
apply (drule star_of_less [THEN iffD2])  | 
|
1877  | 
apply simp  | 
|
1878  | 
apply (blast intro: order_less_trans)  | 
|
1879  | 
done  | 
|
| 27468 | 1880  | 
|
1881  | 
||
1882  | 
lemma Infinitesimal_hypreal_of_nat_iff:  | 
|
| 64435 | 1883  | 
  "Infinitesimal = {x. \<forall>n. hnorm x < inverse (hypreal_of_nat (Suc n))}"
 | 
1884  | 
apply (simp add: Infinitesimal_def)  | 
|
1885  | 
apply (auto simp add: lemma_Infinitesimal2)  | 
|
1886  | 
done  | 
|
| 27468 | 1887  | 
|
1888  | 
||
| 64435 | 1889  | 
subsection \<open>Proof that \<open>\<omega>\<close> is an infinite number\<close>  | 
| 27468 | 1890  | 
|
| 64435 | 1891  | 
text \<open>It will follow that \<open>\<epsilon>\<close> is an infinitesimal number.\<close>  | 
| 27468 | 1892  | 
|
1893  | 
lemma Suc_Un_eq: "{n. n < Suc m} = {n. n < m} Un {n. n = m}"
 | 
|
| 64435 | 1894  | 
by (auto simp add: less_Suc_eq)  | 
| 27468 | 1895  | 
|
| 64435 | 1896  | 
|
| 64438 | 1897  | 
text \<open>Prove that any segment is finite and hence cannot belong to \<open>\<U>\<close>.\<close>  | 
| 27468 | 1898  | 
|
1899  | 
lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}"
 | 
|
| 64435 | 1900  | 
by auto  | 
| 27468 | 1901  | 
|
1902  | 
lemma finite_real_of_nat_less_real: "finite {n::nat. real n < u}"
 | 
|
| 64435 | 1903  | 
apply (cut_tac x = u in reals_Archimedean2, safe)  | 
1904  | 
apply (rule finite_real_of_nat_segment [THEN [2] finite_subset])  | 
|
1905  | 
apply (auto dest: order_less_trans)  | 
|
1906  | 
done  | 
|
| 27468 | 1907  | 
|
| 64435 | 1908  | 
lemma lemma_real_le_Un_eq: "{n. f n \<le> u} = {n. f n < u} \<union> {n. u = (f n :: real)}"
 | 
1909  | 
by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)  | 
|
| 27468 | 1910  | 
|
1911  | 
lemma finite_real_of_nat_le_real: "finite {n::nat. real n \<le> u}"
 | 
|
| 64435 | 1912  | 
by (auto simp add: lemma_real_le_Un_eq lemma_finite_omega_set finite_real_of_nat_less_real)  | 
| 27468 | 1913  | 
|
| 61945 | 1914  | 
lemma finite_rabs_real_of_nat_le_real: "finite {n::nat. \<bar>real n\<bar> \<le> u}"
 | 
| 64435 | 1915  | 
by (simp add: finite_real_of_nat_le_real)  | 
| 27468 | 1916  | 
|
1917  | 
lemma rabs_real_of_nat_le_real_FreeUltrafilterNat:  | 
|
| 64438 | 1918  | 
"\<not> eventually (\<lambda>n. \<bar>real n\<bar> \<le> u) \<U>"  | 
| 64435 | 1919  | 
by (blast intro!: FreeUltrafilterNat.finite finite_rabs_real_of_nat_le_real)  | 
| 27468 | 1920  | 
|
| 64438 | 1921  | 
lemma FreeUltrafilterNat_nat_gt_real: "eventually (\<lambda>n. u < real n) \<U>"  | 
| 64435 | 1922  | 
apply (rule FreeUltrafilterNat.finite')  | 
1923  | 
  apply (subgoal_tac "{n::nat. \<not> u < real n} = {n. real n \<le> u}")
 | 
|
1924  | 
apply (auto simp add: finite_real_of_nat_le_real)  | 
|
1925  | 
done  | 
|
| 27468 | 1926  | 
|
| 64435 | 1927  | 
text \<open>The complement of \<open>{n. \<bar>real n\<bar> \<le> u} = {n. u < \<bar>real n\<bar>}\<close> is in
 | 
| 64438 | 1928  | 
\<open>\<U>\<close> by property of (free) ultrafilters.\<close>  | 
| 27468 | 1929  | 
|
1930  | 
lemma Compl_real_le_eq: "- {n::nat. real n \<le> u} = {n. u < real n}"
 | 
|
| 64435 | 1931  | 
by (auto dest!: order_le_less_trans simp add: linorder_not_le)  | 
| 27468 | 1932  | 
|
| 64435 | 1933  | 
text \<open>@{term \<omega>} is a member of @{term HInfinite}.\<close>
 | 
| 61981 | 1934  | 
theorem HInfinite_omega [simp]: "\<omega> \<in> HInfinite"  | 
| 64435 | 1935  | 
apply (simp add: omega_def)  | 
1936  | 
apply (rule FreeUltrafilterNat_HInfinite)  | 
|
1937  | 
apply clarify  | 
|
1938  | 
apply (rule_tac u1 = "u-1" in eventually_mono [OF FreeUltrafilterNat_nat_gt_real])  | 
|
1939  | 
apply auto  | 
|
1940  | 
done  | 
|
| 27468 | 1941  | 
|
| 64435 | 1942  | 
|
1943  | 
text \<open>Epsilon is a member of Infinitesimal.\<close>  | 
|
| 27468 | 1944  | 
|
| 61981 | 1945  | 
lemma Infinitesimal_epsilon [simp]: "\<epsilon> \<in> Infinitesimal"  | 
| 64435 | 1946  | 
by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega  | 
1947  | 
simp add: hypreal_epsilon_inverse_omega)  | 
|
| 27468 | 1948  | 
|
| 61981 | 1949  | 
lemma HFinite_epsilon [simp]: "\<epsilon> \<in> HFinite"  | 
| 64435 | 1950  | 
by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])  | 
| 27468 | 1951  | 
|
| 61982 | 1952  | 
lemma epsilon_approx_zero [simp]: "\<epsilon> \<approx> 0"  | 
| 64435 | 1953  | 
by (simp add: mem_infmal_iff [symmetric])  | 
| 27468 | 1954  | 
|
| 64435 | 1955  | 
text \<open>Needed for proof that we define a hyperreal \<open>[<X(n)] \<approx> hypreal_of_real a\<close> given  | 
1956  | 
that \<open>\<forall>n. |X n - a| < 1/n\<close>. Used in proof of \<open>NSLIM \<Rightarrow> LIM\<close>.\<close>  | 
|
1957  | 
lemma real_of_nat_less_inverse_iff: "0 < u \<Longrightarrow> u < inverse (real(Suc n)) \<longleftrightarrow> real(Suc n) < inverse u"  | 
|
1958  | 
apply (simp add: inverse_eq_divide)  | 
|
1959  | 
apply (subst pos_less_divide_eq, assumption)  | 
|
1960  | 
apply (subst pos_less_divide_eq)  | 
|
1961  | 
apply simp  | 
|
1962  | 
apply (simp add: mult.commute)  | 
|
1963  | 
done  | 
|
| 27468 | 1964  | 
|
| 64435 | 1965  | 
lemma finite_inverse_real_of_posnat_gt_real: "0 < u \<Longrightarrow> finite {n. u < inverse (real (Suc n))}"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61378 
diff
changeset
 | 
1966  | 
proof (simp only: real_of_nat_less_inverse_iff)  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61378 
diff
changeset
 | 
1967  | 
  have "{n. 1 + real n < inverse u} = {n. real n < inverse u - 1}"
 | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61378 
diff
changeset
 | 
1968  | 
by fastforce  | 
| 64435 | 1969  | 
  then show "finite {n. real (Suc n) < inverse u}"
 | 
1970  | 
using finite_real_of_nat_less_real [of "inverse u - 1"]  | 
|
1971  | 
by auto  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61378 
diff
changeset
 | 
1972  | 
qed  | 
| 27468 | 1973  | 
|
1974  | 
lemma lemma_real_le_Un_eq2:  | 
|
| 64435 | 1975  | 
  "{n. u \<le> inverse(real(Suc n))} =
 | 
1976  | 
    {n. u < inverse(real(Suc n))} \<union> {n. u = inverse(real(Suc n))}"
 | 
|
1977  | 
by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le)  | 
|
| 27468 | 1978  | 
|
| 64435 | 1979  | 
lemma finite_inverse_real_of_posnat_ge_real: "0 < u \<Longrightarrow> finite {n. u \<le> inverse (real (Suc n))}"
 | 
1980  | 
by (auto simp add: lemma_real_le_Un_eq2 lemma_finite_epsilon_set finite_inverse_real_of_posnat_gt_real  | 
|
1981  | 
simp del: of_nat_Suc)  | 
|
| 27468 | 1982  | 
|
1983  | 
lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat:  | 
|
| 64438 | 1984  | 
"0 < u \<Longrightarrow> \<not> eventually (\<lambda>n. u \<le> inverse(real(Suc n))) \<U>"  | 
| 64435 | 1985  | 
by (blast intro!: FreeUltrafilterNat.finite finite_inverse_real_of_posnat_ge_real)  | 
| 27468 | 1986  | 
|
| 64435 | 1987  | 
text \<open>The complement of \<open>{n. u \<le> inverse(real(Suc n))} = {n. inverse (real (Suc n)) < u}\<close>
 | 
| 64438 | 1988  | 
is in \<open>\<U>\<close> by property of (free) ultrafilters.\<close>  | 
| 64435 | 1989  | 
lemma Compl_le_inverse_eq: "- {n. u \<le> inverse(real(Suc n))} = {n. inverse(real(Suc n)) < u}"
 | 
1990  | 
by (auto dest!: order_le_less_trans simp add: linorder_not_le)  | 
|
| 56225 | 1991  | 
|
| 27468 | 1992  | 
|
1993  | 
lemma FreeUltrafilterNat_inverse_real_of_posnat:  | 
|
| 64438 | 1994  | 
"0 < u \<Longrightarrow> eventually (\<lambda>n. inverse(real(Suc n)) < u) \<U>"  | 
| 64435 | 1995  | 
by (drule inverse_real_of_posnat_ge_real_FreeUltrafilterNat)  | 
1996  | 
(simp add: FreeUltrafilterNat.eventually_not_iff not_le[symmetric])  | 
|
| 27468 | 1997  | 
|
| 64435 | 1998  | 
text \<open>Example of an hypersequence (i.e. an extended standard sequence)  | 
1999  | 
whose term with an hypernatural suffix is an infinitesimal i.e.  | 
|
2000  | 
the whn'nth term of the hypersequence is a member of Infinitesimal\<close>  | 
|
| 27468 | 2001  | 
|
| 64435 | 2002  | 
lemma SEQ_Infinitesimal: "( *f* (\<lambda>n::nat. inverse(real(Suc n)))) whn \<in> Infinitesimal"  | 
2003  | 
by (simp add: hypnat_omega_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff  | 
|
2004  | 
FreeUltrafilterNat_inverse_real_of_posnat del: of_nat_Suc)  | 
|
| 27468 | 2005  | 
|
| 64435 | 2006  | 
text \<open>Example where we get a hyperreal from a real sequence  | 
2007  | 
for which a particular property holds. The theorem is  | 
|
2008  | 
used in proofs about equivalence of nonstandard and  | 
|
2009  | 
standard neighbourhoods. Also used for equivalence of  | 
|
2010  | 
nonstandard ans standard definitions of pointwise  | 
|
2011  | 
limit.\<close>  | 
|
| 27468 | 2012  | 
|
| 64435 | 2013  | 
text \<open>\<open>|X(n) - x| < 1/n \<Longrightarrow> [<X n>] - hypreal_of_real x| \<in> Infinitesimal\<close>\<close>  | 
| 27468 | 2014  | 
lemma real_seq_to_hypreal_Infinitesimal:  | 
| 64435 | 2015  | 
"\<forall>n. norm (X n - x) < inverse (real (Suc n)) \<Longrightarrow> star_n X - star_of x \<in> Infinitesimal"  | 
2016  | 
unfolding star_n_diff star_of_def Infinitesimal_FreeUltrafilterNat_iff star_n_inverse  | 
|
2017  | 
by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat  | 
|
2018  | 
intro: order_less_trans elim!: eventually_mono)  | 
|
| 27468 | 2019  | 
|
2020  | 
lemma real_seq_to_hypreal_approx:  | 
|
| 64435 | 2021  | 
"\<forall>n. norm (X n - x) < inverse (real (Suc n)) \<Longrightarrow> star_n X \<approx> star_of x"  | 
2022  | 
by (metis bex_Infinitesimal_iff real_seq_to_hypreal_Infinitesimal)  | 
|
| 27468 | 2023  | 
|
2024  | 
lemma real_seq_to_hypreal_approx2:  | 
|
| 64435 | 2025  | 
"\<forall>n. norm (x - X n) < inverse(real(Suc n)) \<Longrightarrow> star_n X \<approx> star_of x"  | 
2026  | 
by (metis norm_minus_commute real_seq_to_hypreal_approx)  | 
|
| 27468 | 2027  | 
|
2028  | 
lemma real_seq_to_hypreal_Infinitesimal2:  | 
|
| 64435 | 2029  | 
"\<forall>n. norm(X n - Y n) < inverse(real(Suc n)) \<Longrightarrow> star_n X - star_n Y \<in> Infinitesimal"  | 
2030  | 
unfolding Infinitesimal_FreeUltrafilterNat_iff star_n_diff  | 
|
2031  | 
by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat  | 
|
2032  | 
intro: order_less_trans elim!: eventually_mono)  | 
|
| 27468 | 2033  | 
|
2034  | 
end  |