| author | wenzelm | 
| Thu, 01 Sep 2016 14:49:36 +0200 | |
| changeset 63745 | dde79b7faddf | 
| parent 63579 | 73939a9b70a3 | 
| child 64435 | c93b0e6131c3 | 
| permissions | -rw-r--r-- | 
| 62479 | 1 | (* Title: HOL/Nonstandard_Analysis/HLim.thy | 
| 41589 | 2 | Author: Jacques D. Fleuriot, University of Cambridge | 
| 3 | Author: Lawrence C Paulson | |
| 27468 | 4 | *) | 
| 5 | ||
| 61975 | 6 | section\<open>Limits and Continuity (Nonstandard)\<close> | 
| 27468 | 7 | |
| 8 | theory HLim | |
| 63579 | 9 | imports Star | 
| 10 | abbrevs "--->" = "\<midarrow>\<rightarrow>\<^sub>N\<^sub>S" | |
| 27468 | 11 | begin | 
| 12 | ||
| 61975 | 13 | text\<open>Nonstandard Definitions\<close> | 
| 27468 | 14 | |
| 15 | definition | |
| 16 | NSLIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool" | |
| 61971 | 17 |             ("((_)/ \<midarrow>(_)/\<rightarrow>\<^sub>N\<^sub>S (_))" [60, 0, 60] 60) where
 | 
| 18 | "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L = | |
| 61982 | 19 | (\<forall>x. (x \<noteq> star_of a & x \<approx> star_of a --> ( *f* f) x \<approx> star_of L))" | 
| 27468 | 20 | |
| 21 | definition | |
| 22 | isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where | |
| 61975 | 23 | \<comment>\<open>NS definition dispenses with limit notions\<close> | 
| 61982 | 24 | "isNSCont f a = (\<forall>y. y \<approx> star_of a --> | 
| 25 | ( *f* f) y \<approx> star_of (f a))" | |
| 27468 | 26 | |
| 27 | definition | |
| 28 | isNSUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where | |
| 61982 | 29 | "isNSUCont f = (\<forall>x y. x \<approx> y --> ( *f* f) x \<approx> ( *f* f) y)" | 
| 27468 | 30 | |
| 31 | ||
| 61975 | 32 | subsection \<open>Limits of Functions\<close> | 
| 27468 | 33 | |
| 34 | lemma NSLIM_I: | |
| 35 | "(\<And>x. \<lbrakk>x \<noteq> star_of a; x \<approx> star_of a\<rbrakk> \<Longrightarrow> starfun f x \<approx> star_of L) | |
| 61971 | 36 | \<Longrightarrow> f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L" | 
| 27468 | 37 | by (simp add: NSLIM_def) | 
| 38 | ||
| 39 | lemma NSLIM_D: | |
| 61971 | 40 | "\<lbrakk>f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L; x \<noteq> star_of a; x \<approx> star_of a\<rbrakk> | 
| 27468 | 41 | \<Longrightarrow> starfun f x \<approx> star_of L" | 
| 42 | by (simp add: NSLIM_def) | |
| 43 | ||
| 61975 | 44 | text\<open>Proving properties of limits using nonstandard definition. | 
| 45 | The properties hold for standard limits as well!\<close> | |
| 27468 | 46 | |
| 47 | lemma NSLIM_mult: | |
| 48 | fixes l m :: "'a::real_normed_algebra" | |
| 61971 | 49 | shows "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m |] | 
| 50 | ==> (%x. f(x) * g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l * m)" | |
| 27468 | 51 | by (auto simp add: NSLIM_def intro!: approx_mult_HFinite) | 
| 52 | ||
| 53 | lemma starfun_scaleR [simp]: | |
| 54 | "starfun (\<lambda>x. f x *\<^sub>R g x) = (\<lambda>x. scaleHR (starfun f x) (starfun g x))" | |
| 55 | by transfer (rule refl) | |
| 56 | ||
| 57 | lemma NSLIM_scaleR: | |
| 61971 | 58 | "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m |] | 
| 59 | ==> (%x. f(x) *\<^sub>R g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l *\<^sub>R m)" | |
| 27468 | 60 | by (auto simp add: NSLIM_def intro!: approx_scaleR_HFinite) | 
| 61 | ||
| 62 | lemma NSLIM_add: | |
| 61971 | 63 | "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m |] | 
| 64 | ==> (%x. f(x) + g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l + m)" | |
| 27468 | 65 | by (auto simp add: NSLIM_def intro!: approx_add) | 
| 66 | ||
| 61971 | 67 | lemma NSLIM_const [simp]: "(%x. k) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S k" | 
| 27468 | 68 | by (simp add: NSLIM_def) | 
| 69 | ||
| 61971 | 70 | lemma NSLIM_minus: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L ==> (%x. -f(x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S -L" | 
| 27468 | 71 | by (simp add: NSLIM_def) | 
| 72 | ||
| 73 | lemma NSLIM_diff: | |
| 61971 | 74 | "\<lbrakk>f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l - m)" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
51525diff
changeset | 75 | by (simp only: NSLIM_add NSLIM_minus diff_conv_add_uminus) | 
| 27468 | 76 | |
| 61971 | 77 | lemma NSLIM_add_minus: "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S m |] ==> (%x. f(x) + -g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S (l + -m)" | 
| 27468 | 78 | by (simp only: NSLIM_add NSLIM_minus) | 
| 79 | ||
| 80 | lemma NSLIM_inverse: | |
| 81 | fixes L :: "'a::real_normed_div_algebra" | |
| 61971 | 82 | shows "[| f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L; L \<noteq> 0 |] | 
| 83 | ==> (%x. inverse(f(x))) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (inverse L)" | |
| 27468 | 84 | apply (simp add: NSLIM_def, clarify) | 
| 85 | apply (drule spec) | |
| 86 | apply (auto simp add: star_of_approx_inverse) | |
| 87 | done | |
| 88 | ||
| 89 | lemma NSLIM_zero: | |
| 61971 | 90 | assumes f: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S l" shows "(%x. f(x) - l) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S 0" | 
| 27468 | 91 | proof - | 
| 61971 | 92 | have "(\<lambda>x. f x - l) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S l - l" | 
| 27468 | 93 | by (rule NSLIM_diff [OF f NSLIM_const]) | 
| 94 | thus ?thesis by simp | |
| 95 | qed | |
| 96 | ||
| 61971 | 97 | lemma NSLIM_zero_cancel: "(%x. f(x) - l) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0 ==> f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S l" | 
| 27468 | 98 | apply (drule_tac g = "%x. l" and m = l in NSLIM_add) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
54230diff
changeset | 99 | apply (auto simp add: add.assoc) | 
| 27468 | 100 | done | 
| 101 | ||
| 102 | lemma NSLIM_const_not_eq: | |
| 103 | fixes a :: "'a::real_normed_algebra_1" | |
| 61971 | 104 | shows "k \<noteq> L \<Longrightarrow> \<not> (\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L" | 
| 27468 | 105 | apply (simp add: NSLIM_def) | 
| 61981 | 106 | apply (rule_tac x="star_of a + of_hypreal \<epsilon>" in exI) | 
| 27468 | 107 | apply (simp add: hypreal_epsilon_not_zero approx_def) | 
| 108 | done | |
| 109 | ||
| 110 | lemma NSLIM_not_zero: | |
| 111 | fixes a :: "'a::real_normed_algebra_1" | |
| 61971 | 112 | shows "k \<noteq> 0 \<Longrightarrow> \<not> (\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S 0" | 
| 27468 | 113 | by (rule NSLIM_const_not_eq) | 
| 114 | ||
| 115 | lemma NSLIM_const_eq: | |
| 116 | fixes a :: "'a::real_normed_algebra_1" | |
| 61971 | 117 | shows "(\<lambda>x. k) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> k = L" | 
| 27468 | 118 | apply (rule ccontr) | 
| 119 | apply (blast dest: NSLIM_const_not_eq) | |
| 120 | done | |
| 121 | ||
| 122 | lemma NSLIM_unique: | |
| 123 | fixes a :: "'a::real_normed_algebra_1" | |
| 61971 | 124 | shows "\<lbrakk>f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L; f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S M\<rbrakk> \<Longrightarrow> L = M" | 
| 27468 | 125 | apply (drule (1) NSLIM_diff) | 
| 126 | apply (auto dest!: NSLIM_const_eq) | |
| 127 | done | |
| 128 | ||
| 129 | lemma NSLIM_mult_zero: | |
| 130 | fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" | |
| 61971 | 131 | shows "[| f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0; g \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0 |] ==> (%x. f(x)*g(x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0" | 
| 27468 | 132 | by (drule NSLIM_mult, auto) | 
| 133 | ||
| 61971 | 134 | lemma NSLIM_self: "(%x. x) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S a" | 
| 27468 | 135 | by (simp add: NSLIM_def) | 
| 136 | ||
| 61975 | 137 | subsubsection \<open>Equivalence of @{term filterlim} and @{term NSLIM}\<close>
 | 
| 27468 | 138 | |
| 139 | lemma LIM_NSLIM: | |
| 61976 | 140 | assumes f: "f \<midarrow>a\<rightarrow> L" shows "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L" | 
| 27468 | 141 | proof (rule NSLIM_I) | 
| 142 | fix x | |
| 143 | assume neq: "x \<noteq> star_of a" | |
| 144 | assume approx: "x \<approx> star_of a" | |
| 145 | have "starfun f x - star_of L \<in> Infinitesimal" | |
| 146 | proof (rule InfinitesimalI2) | |
| 147 | fix r::real assume r: "0 < r" | |
| 148 | from LIM_D [OF f r] | |
| 149 | obtain s where s: "0 < s" and | |
| 150 | less_r: "\<And>x. \<lbrakk>x \<noteq> a; norm (x - a) < s\<rbrakk> \<Longrightarrow> norm (f x - L) < r" | |
| 151 | by fast | |
| 152 | from less_r have less_r': | |
| 153 | "\<And>x. \<lbrakk>x \<noteq> star_of a; hnorm (x - star_of a) < star_of s\<rbrakk> | |
| 154 | \<Longrightarrow> hnorm (starfun f x - star_of L) < star_of r" | |
| 155 | by transfer | |
| 156 | from approx have "x - star_of a \<in> Infinitesimal" | |
| 157 | by (unfold approx_def) | |
| 158 | hence "hnorm (x - star_of a) < star_of s" | |
| 159 | using s by (rule InfinitesimalD2) | |
| 160 | with neq show "hnorm (starfun f x - star_of L) < star_of r" | |
| 161 | by (rule less_r') | |
| 162 | qed | |
| 163 | thus "starfun f x \<approx> star_of L" | |
| 164 | by (unfold approx_def) | |
| 165 | qed | |
| 166 | ||
| 167 | lemma NSLIM_LIM: | |
| 61976 | 168 | assumes f: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L" shows "f \<midarrow>a\<rightarrow> L" | 
| 27468 | 169 | proof (rule LIM_I) | 
| 170 | fix r::real assume r: "0 < r" | |
| 171 | have "\<exists>s>0. \<forall>x. x \<noteq> star_of a \<and> hnorm (x - star_of a) < s | |
| 172 | \<longrightarrow> hnorm (starfun f x - star_of L) < star_of r" | |
| 173 | proof (rule exI, safe) | |
| 61981 | 174 | show "0 < \<epsilon>" by (rule hypreal_epsilon_gt_zero) | 
| 27468 | 175 | next | 
| 176 | fix x assume neq: "x \<noteq> star_of a" | |
| 61981 | 177 | assume "hnorm (x - star_of a) < \<epsilon>" | 
| 27468 | 178 | with Infinitesimal_epsilon | 
| 179 | have "x - star_of a \<in> Infinitesimal" | |
| 180 | by (rule hnorm_less_Infinitesimal) | |
| 181 | hence "x \<approx> star_of a" | |
| 182 | by (unfold approx_def) | |
| 183 | with f neq have "starfun f x \<approx> star_of L" | |
| 184 | by (rule NSLIM_D) | |
| 185 | hence "starfun f x - star_of L \<in> Infinitesimal" | |
| 186 | by (unfold approx_def) | |
| 187 | thus "hnorm (starfun f x - star_of L) < star_of r" | |
| 188 | using r by (rule InfinitesimalD2) | |
| 189 | qed | |
| 190 | thus "\<exists>s>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < s \<longrightarrow> norm (f x - L) < r" | |
| 191 | by transfer | |
| 192 | qed | |
| 193 | ||
| 61976 | 194 | theorem LIM_NSLIM_iff: "(f \<midarrow>x\<rightarrow> L) = (f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L)" | 
| 27468 | 195 | by (blast intro: LIM_NSLIM NSLIM_LIM) | 
| 196 | ||
| 197 | ||
| 61975 | 198 | subsection \<open>Continuity\<close> | 
| 27468 | 199 | |
| 200 | lemma isNSContD: | |
| 201 | "\<lbrakk>isNSCont f a; y \<approx> star_of a\<rbrakk> \<Longrightarrow> ( *f* f) y \<approx> star_of (f a)" | |
| 202 | by (simp add: isNSCont_def) | |
| 203 | ||
| 61971 | 204 | lemma isNSCont_NSLIM: "isNSCont f a ==> f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a) " | 
| 27468 | 205 | by (simp add: isNSCont_def NSLIM_def) | 
| 206 | ||
| 61971 | 207 | lemma NSLIM_isNSCont: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a) ==> isNSCont f a" | 
| 27468 | 208 | apply (simp add: isNSCont_def NSLIM_def, auto) | 
| 209 | apply (case_tac "y = star_of a", auto) | |
| 210 | done | |
| 211 | ||
| 61975 | 212 | text\<open>NS continuity can be defined using NS Limit in | 
| 63040 | 213 | similar fashion to standard definition of continuity\<close> | 
| 61971 | 214 | lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a))" | 
| 27468 | 215 | by (blast intro: isNSCont_NSLIM NSLIM_isNSCont) | 
| 216 | ||
| 61975 | 217 | text\<open>Hence, NS continuity can be given | 
| 218 | in terms of standard limit\<close> | |
| 61976 | 219 | lemma isNSCont_LIM_iff: "(isNSCont f a) = (f \<midarrow>a\<rightarrow> (f a))" | 
| 27468 | 220 | by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff) | 
| 221 | ||
| 61975 | 222 | text\<open>Moreover, it's trivial now that NS continuity | 
| 223 | is equivalent to standard continuity\<close> | |
| 27468 | 224 | lemma isNSCont_isCont_iff: "(isNSCont f a) = (isCont f a)" | 
| 225 | apply (simp add: isCont_def) | |
| 226 | apply (rule isNSCont_LIM_iff) | |
| 227 | done | |
| 228 | ||
| 61975 | 229 | text\<open>Standard continuity ==> NS continuity\<close> | 
| 27468 | 230 | lemma isCont_isNSCont: "isCont f a ==> isNSCont f a" | 
| 231 | by (erule isNSCont_isCont_iff [THEN iffD2]) | |
| 232 | ||
| 61975 | 233 | text\<open>NS continuity ==> Standard continuity\<close> | 
| 27468 | 234 | lemma isNSCont_isCont: "isNSCont f a ==> isCont f a" | 
| 235 | by (erule isNSCont_isCont_iff [THEN iffD1]) | |
| 236 | ||
| 61975 | 237 | text\<open>Alternative definition of continuity\<close> | 
| 27468 | 238 | |
| 239 | (* Prove equivalence between NS limits - *) | |
| 63040 | 240 | (* seems easier than using standard definition *) | 
| 61971 | 241 | lemma NSLIM_h_iff: "(f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L) = ((%h. f(a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S L)" | 
| 27468 | 242 | apply (simp add: NSLIM_def, auto) | 
| 243 | apply (drule_tac x = "star_of a + x" in spec) | |
| 244 | apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp) | |
| 245 | apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]]) | |
| 246 | apply (erule_tac [3] approx_minus_iff2 [THEN iffD1]) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
54230diff
changeset | 247 | prefer 2 apply (simp add: add.commute) | 
| 27468 | 248 | apply (rule_tac x = x in star_cases) | 
| 249 | apply (rule_tac [2] x = x in star_cases) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
54230diff
changeset | 250 | apply (auto simp add: starfun star_of_def star_n_minus star_n_add add.assoc star_n_zero_num) | 
| 27468 | 251 | done | 
| 252 | ||
| 61971 | 253 | lemma NSLIM_isCont_iff: "(f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S f a) = ((%h. f(a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S f a)" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
51525diff
changeset | 254 | by (fact NSLIM_h_iff) | 
| 27468 | 255 | |
| 256 | lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a" | |
| 257 | by (simp add: isNSCont_def) | |
| 258 | ||
| 259 | lemma isNSCont_inverse: | |
| 260 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_div_algebra" | |
| 261 | shows "[| isNSCont f x; f x \<noteq> 0 |] ==> isNSCont (%x. inverse (f x)) x" | |
| 262 | by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff) | |
| 263 | ||
| 264 | lemma isNSCont_const [simp]: "isNSCont (%x. k) a" | |
| 265 | by (simp add: isNSCont_def) | |
| 266 | ||
| 267 | lemma isNSCont_abs [simp]: "isNSCont abs (a::real)" | |
| 268 | apply (simp add: isNSCont_def) | |
| 269 | apply (auto intro: approx_hrabs simp add: starfun_rabs_hrabs) | |
| 270 | done | |
| 271 | ||
| 272 | ||
| 61975 | 273 | subsection \<open>Uniform Continuity\<close> | 
| 27468 | 274 | |
| 275 | lemma isNSUContD: "[| isNSUCont f; x \<approx> y|] ==> ( *f* f) x \<approx> ( *f* f) y" | |
| 276 | by (simp add: isNSUCont_def) | |
| 277 | ||
| 278 | lemma isUCont_isNSUCont: | |
| 279 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | |
| 280 | assumes f: "isUCont f" shows "isNSUCont f" | |
| 281 | proof (unfold isNSUCont_def, safe) | |
| 282 | fix x y :: "'a star" | |
| 283 | assume approx: "x \<approx> y" | |
| 284 | have "starfun f x - starfun f y \<in> Infinitesimal" | |
| 285 | proof (rule InfinitesimalI2) | |
| 286 | fix r::real assume r: "0 < r" | |
| 287 | with f obtain s where s: "0 < s" and | |
| 288 | less_r: "\<And>x y. norm (x - y) < s \<Longrightarrow> norm (f x - f y) < r" | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
28562diff
changeset | 289 | by (auto simp add: isUCont_def dist_norm) | 
| 27468 | 290 | from less_r have less_r': | 
| 291 | "\<And>x y. hnorm (x - y) < star_of s | |
| 292 | \<Longrightarrow> hnorm (starfun f x - starfun f y) < star_of r" | |
| 293 | by transfer | |
| 294 | from approx have "x - y \<in> Infinitesimal" | |
| 295 | by (unfold approx_def) | |
| 296 | hence "hnorm (x - y) < star_of s" | |
| 297 | using s by (rule InfinitesimalD2) | |
| 298 | thus "hnorm (starfun f x - starfun f y) < star_of r" | |
| 299 | by (rule less_r') | |
| 300 | qed | |
| 301 | thus "starfun f x \<approx> starfun f y" | |
| 302 | by (unfold approx_def) | |
| 303 | qed | |
| 304 | ||
| 305 | lemma isNSUCont_isUCont: | |
| 306 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | |
| 307 | assumes f: "isNSUCont f" shows "isUCont f" | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
28562diff
changeset | 308 | proof (unfold isUCont_def dist_norm, safe) | 
| 27468 | 309 | fix r::real assume r: "0 < r" | 
| 310 | have "\<exists>s>0. \<forall>x y. hnorm (x - y) < s | |
| 311 | \<longrightarrow> hnorm (starfun f x - starfun f y) < star_of r" | |
| 312 | proof (rule exI, safe) | |
| 61981 | 313 | show "0 < \<epsilon>" by (rule hypreal_epsilon_gt_zero) | 
| 27468 | 314 | next | 
| 315 | fix x y :: "'a star" | |
| 61981 | 316 | assume "hnorm (x - y) < \<epsilon>" | 
| 27468 | 317 | with Infinitesimal_epsilon | 
| 318 | have "x - y \<in> Infinitesimal" | |
| 319 | by (rule hnorm_less_Infinitesimal) | |
| 320 | hence "x \<approx> y" | |
| 321 | by (unfold approx_def) | |
| 322 | with f have "starfun f x \<approx> starfun f y" | |
| 323 | by (simp add: isNSUCont_def) | |
| 324 | hence "starfun f x - starfun f y \<in> Infinitesimal" | |
| 325 | by (unfold approx_def) | |
| 326 | thus "hnorm (starfun f x - starfun f y) < star_of r" | |
| 327 | using r by (rule InfinitesimalD2) | |
| 328 | qed | |
| 329 | thus "\<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r" | |
| 330 | by transfer | |
| 331 | qed | |
| 332 | ||
| 333 | end |