| author | huffman | 
| Wed, 19 Mar 2014 20:50:24 -0700 | |
| changeset 56223 | 7696903b9e61 | 
| parent 56217 | dc429a5b13c4 | 
| child 56238 | 5d147e1e18d1 | 
| permissions | -rw-r--r-- | 
| 56215 | 1 | (* Author: John Harrison, Marco Maggesi, Graziano Gentili, Gianni Ciolli, Valentina Bruno | 
| 2 | Ported from "hol_light/Multivariate/canal.ml" by L C Paulson (2014) | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Complex Analysis Basics *}
 | |
| 6 | ||
| 7 | theory Complex_Analysis_Basics | |
| 8 | imports "~~/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space" | |
| 9 | ||
| 10 | begin | |
| 11 | ||
| 12 | subsection {*Complex number lemmas *}
 | |
| 13 | ||
| 14 | lemma abs_sqrt_wlog: | |
| 15 | fixes x::"'a::linordered_idom" | |
| 16 | assumes "!!x::'a. x\<ge>0 \<Longrightarrow> P x (x\<^sup>2)" shows "P \<bar>x\<bar> (x\<^sup>2)" | |
| 17 | by (metis abs_ge_zero assms power2_abs) | |
| 18 | ||
| 19 | lemma complex_abs_le_norm: "abs(Re z) + abs(Im z) \<le> sqrt(2) * norm z" | |
| 20 | proof (cases z) | |
| 21 | case (Complex x y) | |
| 22 | show ?thesis | |
| 23 | apply (rule power2_le_imp_le) | |
| 24 | apply (auto simp: real_sqrt_mult [symmetric] Complex) | |
| 25 | apply (rule abs_sqrt_wlog [where x=x]) | |
| 26 | apply (rule abs_sqrt_wlog [where x=y]) | |
| 27 | apply (simp add: power2_sum add_commute sum_squares_bound) | |
| 28 | done | |
| 29 | qed | |
| 30 | ||
| 31 | lemma continuous_Re: "continuous_on UNIV Re" | |
| 32 | by (metis (poly_guards_query) bounded_linear.continuous_on bounded_linear_Re | |
| 33 | continuous_on_cong continuous_on_id) | |
| 34 | ||
| 35 | lemma continuous_Im: "continuous_on UNIV Im" | |
| 36 | by (metis (poly_guards_query) bounded_linear.continuous_on bounded_linear_Im | |
| 37 | continuous_on_cong continuous_on_id) | |
| 38 | ||
| 39 | lemma open_halfspace_Re_lt: "open {z. Re(z) < b}"
 | |
| 40 | proof - | |
| 41 |   have "{z. Re(z) < b} = Re -`{..<b}"
 | |
| 42 | by blast | |
| 43 | then show ?thesis | |
| 44 | by (auto simp: continuous_Re continuous_imp_open_vimage [of UNIV]) | |
| 45 | qed | |
| 46 | ||
| 47 | lemma open_halfspace_Re_gt: "open {z. Re(z) > b}"
 | |
| 48 | proof - | |
| 49 |   have "{z. Re(z) > b} = Re -`{b<..}"
 | |
| 50 | by blast | |
| 51 | then show ?thesis | |
| 52 | by (auto simp: continuous_Re continuous_imp_open_vimage [of UNIV]) | |
| 53 | qed | |
| 54 | ||
| 55 | lemma closed_halfspace_Re_ge: "closed {z. Re(z) \<ge> b}"
 | |
| 56 | proof - | |
| 57 |   have "{z. Re(z) \<ge> b} = - {z. Re(z) < b}"
 | |
| 58 | by auto | |
| 59 | then show ?thesis | |
| 60 | by (simp add: closed_def open_halfspace_Re_lt) | |
| 61 | qed | |
| 62 | ||
| 63 | lemma closed_halfspace_Re_le: "closed {z. Re(z) \<le> b}"
 | |
| 64 | proof - | |
| 65 |   have "{z. Re(z) \<le> b} = - {z. Re(z) > b}"
 | |
| 66 | by auto | |
| 67 | then show ?thesis | |
| 68 | by (simp add: closed_def open_halfspace_Re_gt) | |
| 69 | qed | |
| 70 | ||
| 71 | lemma closed_halfspace_Re_eq: "closed {z. Re(z) = b}"
 | |
| 72 | proof - | |
| 73 |   have "{z. Re(z) = b} = {z. Re(z) \<le> b} \<inter> {z. Re(z) \<ge> b}"
 | |
| 74 | by auto | |
| 75 | then show ?thesis | |
| 76 | by (auto simp: closed_Int closed_halfspace_Re_le closed_halfspace_Re_ge) | |
| 77 | qed | |
| 78 | ||
| 79 | lemma open_halfspace_Im_lt: "open {z. Im(z) < b}"
 | |
| 80 | proof - | |
| 81 |   have "{z. Im(z) < b} = Im -`{..<b}"
 | |
| 82 | by blast | |
| 83 | then show ?thesis | |
| 84 | by (auto simp: continuous_Im continuous_imp_open_vimage [of UNIV]) | |
| 85 | qed | |
| 86 | ||
| 87 | lemma open_halfspace_Im_gt: "open {z. Im(z) > b}"
 | |
| 88 | proof - | |
| 89 |   have "{z. Im(z) > b} = Im -`{b<..}"
 | |
| 90 | by blast | |
| 91 | then show ?thesis | |
| 92 | by (auto simp: continuous_Im continuous_imp_open_vimage [of UNIV]) | |
| 93 | qed | |
| 94 | ||
| 95 | lemma closed_halfspace_Im_ge: "closed {z. Im(z) \<ge> b}"
 | |
| 96 | proof - | |
| 97 |   have "{z. Im(z) \<ge> b} = - {z. Im(z) < b}"
 | |
| 98 | by auto | |
| 99 | then show ?thesis | |
| 100 | by (simp add: closed_def open_halfspace_Im_lt) | |
| 101 | qed | |
| 102 | ||
| 103 | lemma closed_halfspace_Im_le: "closed {z. Im(z) \<le> b}"
 | |
| 104 | proof - | |
| 105 |   have "{z. Im(z) \<le> b} = - {z. Im(z) > b}"
 | |
| 106 | by auto | |
| 107 | then show ?thesis | |
| 108 | by (simp add: closed_def open_halfspace_Im_gt) | |
| 109 | qed | |
| 110 | ||
| 111 | lemma closed_halfspace_Im_eq: "closed {z. Im(z) = b}"
 | |
| 112 | proof - | |
| 113 |   have "{z. Im(z) = b} = {z. Im(z) \<le> b} \<inter> {z. Im(z) \<ge> b}"
 | |
| 114 | by auto | |
| 115 | then show ?thesis | |
| 116 | by (auto simp: closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge) | |
| 117 | qed | |
| 118 | ||
| 119 | lemma complex_is_Real_iff: "z \<in> \<real> \<longleftrightarrow> Im z = 0" | |
| 120 | by (metis Complex_in_Reals Im_complex_of_real Reals_cases complex_surj) | |
| 121 | ||
| 122 | lemma closed_complex_Reals: "closed (Reals :: complex set)" | |
| 123 | proof - | |
| 124 |   have "-(Reals :: complex set) = {z. Im(z) < 0} \<union> {z. 0 < Im(z)}"
 | |
| 125 | by (auto simp: complex_is_Real_iff) | |
| 126 | then show ?thesis | |
| 127 | by (metis closed_def open_Un open_halfspace_Im_gt open_halfspace_Im_lt) | |
| 128 | qed | |
| 129 | ||
| 130 | ||
| 131 | lemma linear_times: | |
| 132 |   fixes c::"'a::{real_algebra,real_vector}" shows "linear (\<lambda>x. c * x)"
 | |
| 133 | by (auto simp: linearI distrib_left) | |
| 134 | ||
| 135 | lemma bilinear_times: | |
| 136 |   fixes c::"'a::{real_algebra,real_vector}" shows "bilinear (\<lambda>x y::'a. x*y)"
 | |
| 137 | unfolding bilinear_def | |
| 138 | by (auto simp: distrib_left distrib_right intro!: linearI) | |
| 139 | ||
| 140 | lemma linear_cnj: "linear cnj" | |
| 141 | by (auto simp: linearI cnj_def) | |
| 142 | ||
| 143 | lemma tendsto_mult_left: | |
| 144 | fixes c::"'a::real_normed_algebra" | |
| 145 | shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) ---> c * l) F" | |
| 146 | by (rule tendsto_mult [OF tendsto_const]) | |
| 147 | ||
| 148 | lemma tendsto_mult_right: | |
| 149 | fixes c::"'a::real_normed_algebra" | |
| 150 | shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) ---> l * c) F" | |
| 151 | by (rule tendsto_mult [OF _ tendsto_const]) | |
| 152 | ||
| 153 | lemma tendsto_Re_upper: | |
| 154 | assumes "~ (trivial_limit F)" | |
| 155 | "(f ---> l) F" | |
| 156 | "eventually (\<lambda>x. Re(f x) \<le> b) F" | |
| 157 | shows "Re(l) \<le> b" | |
| 158 | by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Re) | |
| 159 | ||
| 160 | lemma tendsto_Re_lower: | |
| 161 | assumes "~ (trivial_limit F)" | |
| 162 | "(f ---> l) F" | |
| 163 | "eventually (\<lambda>x. b \<le> Re(f x)) F" | |
| 164 | shows "b \<le> Re(l)" | |
| 165 | by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Re) | |
| 166 | ||
| 167 | lemma tendsto_Im_upper: | |
| 168 | assumes "~ (trivial_limit F)" | |
| 169 | "(f ---> l) F" | |
| 170 | "eventually (\<lambda>x. Im(f x) \<le> b) F" | |
| 171 | shows "Im(l) \<le> b" | |
| 172 | by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Im) | |
| 173 | ||
| 174 | lemma tendsto_Im_lower: | |
| 175 | assumes "~ (trivial_limit F)" | |
| 176 | "(f ---> l) F" | |
| 177 | "eventually (\<lambda>x. b \<le> Im(f x)) F" | |
| 178 | shows "b \<le> Im(l)" | |
| 179 | by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Im) | |
| 180 | ||
| 181 | subsection{*General lemmas*}
 | |
| 182 | ||
| 183 | lemma continuous_mult_left: | |
| 184 | fixes c::"'a::real_normed_algebra" | |
| 185 | shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. c * f x)" | |
| 186 | by (rule continuous_mult [OF continuous_const]) | |
| 187 | ||
| 188 | lemma continuous_mult_right: | |
| 189 | fixes c::"'a::real_normed_algebra" | |
| 190 | shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x * c)" | |
| 191 | by (rule continuous_mult [OF _ continuous_const]) | |
| 192 | ||
| 193 | lemma continuous_on_mult_left: | |
| 194 | fixes c::"'a::real_normed_algebra" | |
| 195 | shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. c * f x)" | |
| 196 | by (rule continuous_on_mult [OF continuous_on_const]) | |
| 197 | ||
| 198 | lemma continuous_on_mult_right: | |
| 199 | fixes c::"'a::real_normed_algebra" | |
| 200 | shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x * c)" | |
| 201 | by (rule continuous_on_mult [OF _ continuous_on_const]) | |
| 202 | ||
| 203 | lemma uniformly_continuous_on_cmul_right [continuous_on_intros]: | |
| 204 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" | |
| 205 | assumes "uniformly_continuous_on s f" | |
| 206 | shows "uniformly_continuous_on s (\<lambda>x. f x * c)" | |
| 207 | proof (cases "c=0") | |
| 208 | case True then show ?thesis | |
| 209 | by (simp add: uniformly_continuous_on_const) | |
| 210 | next | |
| 211 | case False show ?thesis | |
| 212 | apply (rule bounded_linear.uniformly_continuous_on) | |
| 213 | apply (metis bounded_linear_ident) | |
| 214 | using assms | |
| 215 | apply (auto simp: uniformly_continuous_on_def dist_norm) | |
| 216 | apply (drule_tac x = "e / norm c" in spec, auto) | |
| 217 | apply (metis divide_pos_pos zero_less_norm_iff False) | |
| 218 | apply (rule_tac x=d in exI, auto) | |
| 219 | apply (drule_tac x = x in bspec, assumption) | |
| 220 | apply (drule_tac x = "x'" in bspec) | |
| 221 | apply (auto simp: False less_divide_eq) | |
| 222 | by (metis dual_order.strict_trans2 left_diff_distrib norm_mult_ineq) | |
| 223 | qed | |
| 224 | ||
| 225 | lemma uniformly_continuous_on_cmul_left[continuous_on_intros]: | |
| 226 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" | |
| 227 | assumes "uniformly_continuous_on s f" | |
| 228 | shows "uniformly_continuous_on s (\<lambda>x. c * f x)" | |
| 229 | by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right) | |
| 230 | ||
| 231 | lemma continuous_within_norm_id [continuous_intros]: "continuous (at x within S) norm" | |
| 232 | by (rule continuous_norm [OF continuous_ident]) | |
| 233 | ||
| 234 | lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm" | |
| 235 | by (metis continuous_on_eq continuous_on_id continuous_on_norm) | |
| 236 | ||
| 237 | ||
| 238 | subsection{*DERIV stuff*}
 | |
| 239 | ||
| 240 | (*move some premises to a sensible order. Use more \<And> symbols.*) | |
| 241 | ||
| 242 | lemma DERIV_continuous_on: "\<lbrakk>\<And>x. x \<in> s \<Longrightarrow> DERIV f x :> D\<rbrakk> \<Longrightarrow> continuous_on s f" | |
| 243 | by (metis DERIV_continuous continuous_at_imp_continuous_on) | |
| 244 | ||
| 245 | lemma DERIV_subset: | |
| 246 | "(f has_field_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s | |
| 247 | \<Longrightarrow> (f has_field_derivative f') (at x within t)" | |
| 248 | by (simp add: has_field_derivative_def has_derivative_within_subset) | |
| 249 | ||
| 250 | lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = op * 0" | |
| 251 | by auto | |
| 252 | ||
| 253 | lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = op * 1" | |
| 254 | by auto | |
| 255 | ||
| 256 | lemma has_derivative_zero_constant: | |
| 257 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 258 | assumes "convex s" | |
| 259 | and d0: "\<And>x. x\<in>s \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within s)" | |
| 260 | shows "\<exists>c. \<forall>x\<in>s. f x = c" | |
| 261 | proof (cases "s={}")
 | |
| 262 | case False | |
| 263 | then obtain x where "x \<in> s" | |
| 264 | by auto | |
| 265 | have d0': "\<forall>x\<in>s. (f has_derivative (\<lambda>h. 0)) (at x within s)" | |
| 266 | by (metis d0) | |
| 267 | have "\<And>y. y \<in> s \<Longrightarrow> f x = f y" | |
| 268 | proof - | |
| 269 | case goal1 | |
| 270 | then show ?case | |
| 271 | using differentiable_bound[OF assms(1) d0', of 0 x y] and `x \<in> s` | |
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 272 | unfolding onorm_zero | 
| 56215 | 273 | by auto | 
| 274 | qed | |
| 275 | then show ?thesis | |
| 276 | by metis | |
| 277 | next | |
| 278 | case True | |
| 279 | then show ?thesis by auto | |
| 280 | qed | |
| 281 | ||
| 282 | lemma has_derivative_zero_unique: | |
| 283 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 284 | assumes "convex s" | |
| 285 | and "\<And>x. x\<in>s \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within s)" | |
| 286 | and "a \<in> s" | |
| 287 | and "x \<in> s" | |
| 288 | shows "f x = f a" | |
| 289 | using assms | |
| 290 | by (metis has_derivative_zero_constant) | |
| 291 | ||
| 292 | lemma has_derivative_zero_connected_constant: | |
| 293 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | |
| 294 | assumes "connected s" | |
| 295 | and "open s" | |
| 296 | and "finite k" | |
| 297 | and "continuous_on s f" | |
| 298 | and "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)" | |
| 299 | obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c" | |
| 300 | proof (cases "s = {}")
 | |
| 301 | case True | |
| 302 | then show ?thesis | |
| 303 | by (metis empty_iff that) | |
| 304 | next | |
| 305 | case False | |
| 306 | then obtain c where "c \<in> s" | |
| 307 | by (metis equals0I) | |
| 308 | then show ?thesis | |
| 309 | by (metis has_derivative_zero_unique_strong_connected assms that) | |
| 310 | qed | |
| 311 | ||
| 312 | lemma DERIV_zero_connected_constant: | |
| 313 |   fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
 | |
| 314 | assumes "connected s" | |
| 315 | and "open s" | |
| 316 | and "finite k" | |
| 317 | and "continuous_on s f" | |
| 318 | and "\<forall>x\<in>(s - k). DERIV f x :> 0" | |
| 319 | obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c" | |
| 320 | using has_derivative_zero_connected_constant [OF assms(1-4)] assms | |
| 321 | by (metis DERIV_const Derivative.has_derivative_const Diff_iff at_within_open | |
| 322 | frechet_derivative_at has_field_derivative_def) | |
| 323 | ||
| 324 | lemma DERIV_zero_constant: | |
| 325 |   fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
 | |
| 326 | shows "\<lbrakk>convex s; | |
| 327 | \<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)\<rbrakk> | |
| 328 | \<Longrightarrow> \<exists>c. \<forall>x \<in> s. f(x) = c" | |
| 329 | unfolding has_field_derivative_def | |
| 330 | by (auto simp: lambda_zero intro: has_derivative_zero_constant) | |
| 331 | ||
| 332 | lemma DERIV_zero_unique: | |
| 333 |   fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
 | |
| 334 | assumes "convex s" | |
| 335 | and d0: "\<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)" | |
| 336 | and "a \<in> s" | |
| 337 | and "x \<in> s" | |
| 338 | shows "f x = f a" | |
| 339 | apply (rule has_derivative_zero_unique [where f=f, OF assms(1) _ assms(3,4)]) | |
| 340 | by (metis d0 has_field_derivative_imp_has_derivative lambda_zero) | |
| 341 | ||
| 342 | lemma DERIV_zero_connected_unique: | |
| 343 |   fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
 | |
| 344 | assumes "connected s" | |
| 345 | and "open s" | |
| 346 | and d0: "\<And>x. x\<in>s \<Longrightarrow> DERIV f x :> 0" | |
| 347 | and "a \<in> s" | |
| 348 | and "x \<in> s" | |
| 349 | shows "f x = f a" | |
| 350 |     apply (rule Integration.has_derivative_zero_unique_strong_connected [of s "{}" f])
 | |
| 351 | using assms | |
| 352 | apply auto | |
| 353 | apply (metis DERIV_continuous_on) | |
| 354 | by (metis at_within_open has_field_derivative_def lambda_zero) | |
| 355 | ||
| 356 | lemma DERIV_transform_within: | |
| 357 | assumes "(f has_field_derivative f') (at a within s)" | |
| 358 | and "0 < d" "a \<in> s" | |
| 359 | and "\<And>x. x\<in>s \<Longrightarrow> dist x a < d \<Longrightarrow> f x = g x" | |
| 360 | shows "(g has_field_derivative f') (at a within s)" | |
| 361 | using assms unfolding has_field_derivative_def | |
| 362 | by (blast intro: Derivative.has_derivative_transform_within) | |
| 363 | ||
| 364 | lemma DERIV_transform_within_open: | |
| 365 | assumes "DERIV f a :> f'" | |
| 366 | and "open s" "a \<in> s" | |
| 367 | and "\<And>x. x\<in>s \<Longrightarrow> f x = g x" | |
| 368 | shows "DERIV g a :> f'" | |
| 369 | using assms unfolding has_field_derivative_def | |
| 370 | by (metis has_derivative_transform_within_open) | |
| 371 | ||
| 372 | lemma DERIV_transform_at: | |
| 373 | assumes "DERIV f a :> f'" | |
| 374 | and "0 < d" | |
| 375 | and "\<And>x. dist x a < d \<Longrightarrow> f x = g x" | |
| 376 | shows "DERIV g a :> f'" | |
| 377 | by (blast intro: assms DERIV_transform_within) | |
| 378 | ||
| 379 | ||
| 380 | subsection{*Holomorphic functions*}
 | |
| 381 | ||
| 382 | lemma has_derivative_ident[has_derivative_intros, simp]: | |
| 383 | "FDERIV complex_of_real x :> complex_of_real" | |
| 384 | by (simp add: has_derivative_def tendsto_const bounded_linear_of_real) | |
| 385 | ||
| 386 | lemma has_real_derivative: | |
| 387 | fixes f :: "real\<Rightarrow>real" | |
| 388 | assumes "(f has_derivative f') F" | |
| 389 | obtains c where "(f has_derivative (\<lambda>x. x * c)) F" | |
| 390 | proof - | |
| 391 | obtain c where "f' = (\<lambda>x. x * c)" | |
| 392 | by (metis assms derivative_linear real_bounded_linear) | |
| 393 | then show ?thesis | |
| 394 | by (metis assms that) | |
| 395 | qed | |
| 396 | ||
| 397 | lemma has_real_derivative_iff: | |
| 398 | fixes f :: "real\<Rightarrow>real" | |
| 399 | shows "(\<exists>f'. (f has_derivative (\<lambda>x. x * f')) F) = (\<exists>D. (f has_derivative D) F)" | |
| 400 | by (auto elim: has_real_derivative) | |
| 401 | ||
| 402 | definition complex_differentiable :: "[complex \<Rightarrow> complex, complex filter] \<Rightarrow> bool" | |
| 403 | (infixr "(complex'_differentiable)" 50) | |
| 404 | where "f complex_differentiable F \<equiv> \<exists>f'. (f has_field_derivative f') F" | |
| 405 | ||
| 406 | definition DD :: "['a \<Rightarrow> 'a::real_normed_field, 'a] \<Rightarrow> 'a" --{*for real, complex?*}
 | |
| 407 | where "DD f x \<equiv> THE f'. (f has_derivative (\<lambda>x. x * f')) (at x)" | |
| 408 | ||
| 409 | definition holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool" | |
| 410 | (infixl "(holomorphic'_on)" 50) | |
| 411 | where "f holomorphic_on s \<equiv> \<forall>x \<in> s. \<exists>f'. (f has_field_derivative f') (at x within s)" | |
| 412 | ||
| 413 | lemma holomorphic_on_empty: "f holomorphic_on {}"
 | |
| 414 | by (simp add: holomorphic_on_def) | |
| 415 | ||
| 416 | lemma holomorphic_on_differentiable: | |
| 417 | "f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. f complex_differentiable (at x within s))" | |
| 418 | unfolding holomorphic_on_def complex_differentiable_def has_field_derivative_def | |
| 419 | by (metis mult_commute_abs) | |
| 420 | ||
| 421 | lemma holomorphic_on_open: | |
| 422 | "open s \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>f'. DERIV f x :> f')" | |
| 423 | by (auto simp: holomorphic_on_def has_field_derivative_def at_within_open [of _ s]) | |
| 424 | ||
| 425 | lemma complex_differentiable_imp_continuous_at: | |
| 426 | "f complex_differentiable (at x) \<Longrightarrow> continuous (at x) f" | |
| 427 | by (metis DERIV_continuous complex_differentiable_def) | |
| 428 | ||
| 429 | lemma holomorphic_on_imp_continuous_on: | |
| 430 | "f holomorphic_on s \<Longrightarrow> continuous_on s f" | |
| 431 | by (metis DERIV_continuous continuous_on_eq_continuous_within holomorphic_on_def) | |
| 432 | ||
| 433 | lemma has_derivative_within_open: | |
| 434 | "a \<in> s \<Longrightarrow> open s \<Longrightarrow> (f has_field_derivative f') (at a within s) \<longleftrightarrow> DERIV f a :> f'" | |
| 435 | by (simp add: has_field_derivative_def) (metis has_derivative_within_open) | |
| 436 | ||
| 437 | lemma holomorphic_on_subset: | |
| 438 | "f holomorphic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f holomorphic_on t" | |
| 439 | unfolding holomorphic_on_def | |
| 440 | by (metis DERIV_subset subsetD) | |
| 441 | ||
| 442 | lemma complex_differentiable_within_subset: | |
| 443 | "\<lbrakk>f complex_differentiable (at x within s); t \<subseteq> s\<rbrakk> | |
| 444 | \<Longrightarrow> f complex_differentiable (at x within t)" | |
| 445 | unfolding complex_differentiable_def | |
| 446 | by (metis DERIV_subset) | |
| 447 | ||
| 448 | lemma complex_differentiable_at_within: | |
| 449 | "\<lbrakk>f complex_differentiable (at x)\<rbrakk> | |
| 450 | \<Longrightarrow> f complex_differentiable (at x within s)" | |
| 451 | unfolding complex_differentiable_def | |
| 452 | by (metis DERIV_subset top_greatest) | |
| 453 | ||
| 454 | ||
| 455 | lemma has_derivative_mult_right: | |
| 456 | fixes c:: "'a :: real_normed_algebra" | |
| 457 | shows "((op * c) has_derivative (op * c)) F" | |
| 458 | by (rule has_derivative_mult_right [OF has_derivative_id]) | |
| 459 | ||
| 460 | lemma complex_differentiable_linear: | |
| 461 | "(op * c) complex_differentiable F" | |
| 462 | proof - | |
| 463 | have "\<And>u::complex. (\<lambda>x. x * u) = op * u" | |
| 464 | by (rule ext) (simp add: mult_ac) | |
| 465 | then show ?thesis | |
| 466 | unfolding complex_differentiable_def has_field_derivative_def | |
| 467 | by (force intro: has_derivative_mult_right) | |
| 468 | qed | |
| 469 | ||
| 470 | lemma complex_differentiable_const: | |
| 471 | "(\<lambda>z. c) complex_differentiable F" | |
| 472 | unfolding complex_differentiable_def has_field_derivative_def | |
| 473 | apply (rule exI [where x=0]) | |
| 474 | by (metis Derivative.has_derivative_const lambda_zero) | |
| 475 | ||
| 476 | lemma complex_differentiable_id: | |
| 477 | "(\<lambda>z. z) complex_differentiable F" | |
| 478 | unfolding complex_differentiable_def has_field_derivative_def | |
| 479 | apply (rule exI [where x=1]) | |
| 480 | apply (simp add: lambda_one [symmetric]) | |
| 481 | done | |
| 482 | ||
| 483 | (*DERIV_minus*) | |
| 484 | lemma field_differentiable_minus: | |
| 485 | assumes "(f has_field_derivative f') F" | |
| 486 | shows "((\<lambda>z. - (f z)) has_field_derivative -f') F" | |
| 487 | apply (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus]) | |
| 488 | using assms | |
| 489 | by (auto simp: has_field_derivative_def field_simps mult_commute_abs) | |
| 490 | ||
| 491 | (*DERIV_add*) | |
| 492 | lemma field_differentiable_add: | |
| 493 | assumes "(f has_field_derivative f') F" "(g has_field_derivative g') F" | |
| 494 | shows "((\<lambda>z. f z + g z) has_field_derivative f' + g') F" | |
| 495 | apply (rule has_derivative_imp_has_field_derivative[OF has_derivative_add]) | |
| 496 | using assms | |
| 497 | by (auto simp: has_field_derivative_def field_simps mult_commute_abs) | |
| 498 | ||
| 499 | (*DERIV_diff*) | |
| 500 | lemma field_differentiable_diff: | |
| 501 | assumes "(f has_field_derivative f') F" "(g has_field_derivative g') F" | |
| 502 | shows "((\<lambda>z. f z - g z) has_field_derivative f' - g') F" | |
| 503 | by (simp only: assms diff_conv_add_uminus field_differentiable_add field_differentiable_minus) | |
| 504 | ||
| 505 | lemma complex_differentiable_minus: | |
| 506 | "f complex_differentiable F \<Longrightarrow> (\<lambda>z. -(f z)) complex_differentiable F" | |
| 507 | using assms unfolding complex_differentiable_def | |
| 508 | by (metis field_differentiable_minus) | |
| 509 | ||
| 510 | lemma complex_differentiable_add: | |
| 511 | assumes "f complex_differentiable F" "g complex_differentiable F" | |
| 512 | shows "(\<lambda>z. f z + g z) complex_differentiable F" | |
| 513 | using assms unfolding complex_differentiable_def | |
| 514 | by (metis field_differentiable_add) | |
| 515 | ||
| 516 | lemma complex_differentiable_diff: | |
| 517 | assumes "f complex_differentiable F" "g complex_differentiable F" | |
| 518 | shows "(\<lambda>z. f z - g z) complex_differentiable F" | |
| 519 | using assms unfolding complex_differentiable_def | |
| 520 | by (metis field_differentiable_diff) | |
| 521 | ||
| 522 | lemma complex_differentiable_inverse: | |
| 523 | assumes "f complex_differentiable (at a within s)" "f a \<noteq> 0" | |
| 524 | shows "(\<lambda>z. inverse (f z)) complex_differentiable (at a within s)" | |
| 525 | using assms unfolding complex_differentiable_def | |
| 526 | by (metis DERIV_inverse_fun) | |
| 527 | ||
| 528 | lemma complex_differentiable_mult: | |
| 529 | assumes "f complex_differentiable (at a within s)" | |
| 530 | "g complex_differentiable (at a within s)" | |
| 531 | shows "(\<lambda>z. f z * g z) complex_differentiable (at a within s)" | |
| 532 | using assms unfolding complex_differentiable_def | |
| 533 | by (metis DERIV_mult [of f _ a s g]) | |
| 534 | ||
| 535 | lemma complex_differentiable_divide: | |
| 536 | assumes "f complex_differentiable (at a within s)" | |
| 537 | "g complex_differentiable (at a within s)" | |
| 538 | "g a \<noteq> 0" | |
| 539 | shows "(\<lambda>z. f z / g z) complex_differentiable (at a within s)" | |
| 540 | using assms unfolding complex_differentiable_def | |
| 541 | by (metis DERIV_divide [of f _ a s g]) | |
| 542 | ||
| 543 | lemma complex_differentiable_power: | |
| 544 | assumes "f complex_differentiable (at a within s)" | |
| 545 | shows "(\<lambda>z. f z ^ n) complex_differentiable (at a within s)" | |
| 546 | using assms unfolding complex_differentiable_def | |
| 547 | by (metis DERIV_power) | |
| 548 | ||
| 549 | lemma complex_differentiable_transform_within: | |
| 550 | "0 < d \<Longrightarrow> | |
| 551 | x \<in> s \<Longrightarrow> | |
| 552 | (\<And>x'. x' \<in> s \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x') \<Longrightarrow> | |
| 553 | f complex_differentiable (at x within s) | |
| 554 | \<Longrightarrow> g complex_differentiable (at x within s)" | |
| 555 | unfolding complex_differentiable_def has_field_derivative_def | |
| 556 | by (blast intro: has_derivative_transform_within) | |
| 557 | ||
| 558 | lemma complex_differentiable_compose_within: | |
| 559 | assumes "f complex_differentiable (at a within s)" | |
| 560 | "g complex_differentiable (at (f a) within f`s)" | |
| 561 | shows "(g o f) complex_differentiable (at a within s)" | |
| 562 | using assms unfolding complex_differentiable_def | |
| 563 | by (metis DERIV_image_chain) | |
| 564 | ||
| 565 | lemma complex_differentiable_within_open: | |
| 566 | "\<lbrakk>a \<in> s; open s\<rbrakk> \<Longrightarrow> f complex_differentiable at a within s \<longleftrightarrow> | |
| 567 | f complex_differentiable at a" | |
| 568 | unfolding complex_differentiable_def | |
| 569 | by (metis at_within_open) | |
| 570 | ||
| 571 | lemma holomorphic_transform: | |
| 572 | "\<lbrakk>f holomorphic_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g holomorphic_on s" | |
| 573 | apply (auto simp: holomorphic_on_def has_field_derivative_def) | |
| 574 | by (metis complex_differentiable_def complex_differentiable_transform_within has_field_derivative_def linordered_field_no_ub) | |
| 575 | ||
| 576 | lemma holomorphic_eq: | |
| 577 | "(\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> g holomorphic_on s" | |
| 578 | by (metis holomorphic_transform) | |
| 579 | ||
| 580 | subsection{*Holomorphic*}
 | |
| 581 | ||
| 582 | lemma holomorphic_on_linear: | |
| 583 | "(op * c) holomorphic_on s" | |
| 584 | unfolding holomorphic_on_def by (metis DERIV_cmult_Id) | |
| 585 | ||
| 586 | lemma holomorphic_on_const: | |
| 587 | "(\<lambda>z. c) holomorphic_on s" | |
| 588 | unfolding holomorphic_on_def | |
| 589 | by (metis DERIV_const) | |
| 590 | ||
| 591 | lemma holomorphic_on_id: | |
| 592 | "id holomorphic_on s" | |
| 593 | unfolding holomorphic_on_def id_def | |
| 594 | by (metis DERIV_ident) | |
| 595 | ||
| 596 | lemma holomorphic_on_compose: | |
| 597 | "f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s) | |
| 598 | \<Longrightarrow> (g o f) holomorphic_on s" | |
| 599 | unfolding holomorphic_on_def | |
| 600 | by (metis DERIV_image_chain imageI) | |
| 601 | ||
| 602 | lemma holomorphic_on_compose_gen: | |
| 603 | "\<lbrakk>f holomorphic_on s; g holomorphic_on t; f ` s \<subseteq> t\<rbrakk> \<Longrightarrow> (g o f) holomorphic_on s" | |
| 604 | unfolding holomorphic_on_def | |
| 605 | by (metis DERIV_image_chain DERIV_subset image_subset_iff) | |
| 606 | ||
| 607 | lemma holomorphic_on_minus: | |
| 608 | "f holomorphic_on s \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on s" | |
| 609 | unfolding holomorphic_on_def | |
| 610 | by (metis DERIV_minus) | |
| 611 | ||
| 612 | lemma holomorphic_on_add: | |
| 613 | "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) holomorphic_on s" | |
| 614 | unfolding holomorphic_on_def | |
| 615 | by (metis DERIV_add) | |
| 616 | ||
| 617 | lemma holomorphic_on_diff: | |
| 618 | "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) holomorphic_on s" | |
| 619 | unfolding holomorphic_on_def | |
| 620 | by (metis DERIV_diff) | |
| 621 | ||
| 622 | lemma holomorphic_on_mult: | |
| 623 | "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) holomorphic_on s" | |
| 624 | unfolding holomorphic_on_def | |
| 625 | by auto (metis DERIV_mult) | |
| 626 | ||
| 627 | lemma holomorphic_on_inverse: | |
| 628 | "\<lbrakk>f holomorphic_on s; \<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. inverse (f z)) holomorphic_on s" | |
| 629 | unfolding holomorphic_on_def | |
| 630 | by (metis DERIV_inverse') | |
| 631 | ||
| 632 | lemma holomorphic_on_divide: | |
| 633 | "\<lbrakk>f holomorphic_on s; g holomorphic_on s; \<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. f z / g z) holomorphic_on s" | |
| 634 | unfolding holomorphic_on_def | |
| 635 | by (metis (full_types) DERIV_divide) | |
| 636 | ||
| 637 | lemma holomorphic_on_power: | |
| 638 | "f holomorphic_on s \<Longrightarrow> (\<lambda>z. (f z)^n) holomorphic_on s" | |
| 639 | unfolding holomorphic_on_def | |
| 640 | by (metis DERIV_power) | |
| 641 | ||
| 642 | lemma holomorphic_on_setsum: | |
| 643 | "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) | |
| 644 | \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) holomorphic_on s" | |
| 645 | unfolding holomorphic_on_def | |
| 646 | apply (induct I rule: finite_induct) | |
| 647 | apply (force intro: DERIV_const DERIV_add)+ | |
| 648 | done | |
| 649 | ||
| 650 | lemma DERIV_imp_DD: "DERIV f x :> f' \<Longrightarrow> DD f x = f'" | |
| 651 | apply (simp add: DD_def has_field_derivative_def mult_commute_abs) | |
| 652 | apply (rule the_equality, assumption) | |
| 653 | apply (metis DERIV_unique has_field_derivative_def) | |
| 654 | done | |
| 655 | ||
| 656 | lemma DD_iff_derivative_differentiable: | |
| 657 | fixes f :: "real\<Rightarrow>real" | |
| 658 | shows "DERIV f x :> DD f x \<longleftrightarrow> f differentiable at x" | |
| 659 | unfolding DD_def differentiable_def | |
| 660 | by (metis (full_types) DD_def DERIV_imp_DD has_field_derivative_def has_real_derivative_iff | |
| 661 | mult_commute_abs) | |
| 662 | ||
| 663 | lemma DD_iff_derivative_complex_differentiable: | |
| 664 | fixes f :: "complex\<Rightarrow>complex" | |
| 665 | shows "DERIV f x :> DD f x \<longleftrightarrow> f complex_differentiable at x" | |
| 666 | unfolding DD_def complex_differentiable_def | |
| 667 | by (metis DD_def DERIV_imp_DD) | |
| 668 | ||
| 669 | lemma complex_differentiable_compose: | |
| 670 | "f complex_differentiable at z \<Longrightarrow> g complex_differentiable at (f z) | |
| 671 | \<Longrightarrow> (g o f) complex_differentiable at z" | |
| 672 | by (metis complex_differentiable_at_within complex_differentiable_compose_within) | |
| 673 | ||
| 674 | lemma complex_derivative_chain: | |
| 675 | fixes z::complex | |
| 676 | shows | |
| 677 | "f complex_differentiable at z \<Longrightarrow> g complex_differentiable at (f z) | |
| 678 | \<Longrightarrow> DD (g o f) z = DD g (f z) * DD f z" | |
| 679 | by (metis DD_iff_derivative_complex_differentiable DERIV_chain DERIV_imp_DD) | |
| 680 | ||
| 681 | lemma comp_derivative_chain: | |
| 682 | fixes z::real | |
| 683 | shows "\<lbrakk>f differentiable at z; g differentiable at (f z)\<rbrakk> | |
| 684 | \<Longrightarrow> DD (g o f) z = DD g (f z) * DD f z" | |
| 685 | by (metis DD_iff_derivative_differentiable DERIV_chain DERIV_imp_DD) | |
| 686 | ||
| 687 | lemma complex_derivative_linear: "DD (\<lambda>w. c * w) = (\<lambda>z. c)" | |
| 688 | by (metis DERIV_imp_DD DERIV_cmult_Id) | |
| 689 | ||
| 690 | lemma complex_derivative_ident: "DD (\<lambda>w. w) = (\<lambda>z. 1)" | |
| 691 | by (metis DERIV_imp_DD DERIV_ident) | |
| 692 | ||
| 693 | lemma complex_derivative_const: "DD (\<lambda>w. c) = (\<lambda>z. 0)" | |
| 694 | by (metis DERIV_imp_DD DERIV_const) | |
| 695 | ||
| 696 | lemma complex_derivative_add: | |
| 697 | "\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk> | |
| 698 | \<Longrightarrow> DD (\<lambda>w. f w + g w) z = DD f z + DD g z" | |
| 699 | unfolding complex_differentiable_def | |
| 700 | by (rule DERIV_imp_DD) (metis (poly_guards_query) DERIV_add DERIV_imp_DD) | |
| 701 | ||
| 702 | lemma complex_derivative_diff: | |
| 703 | "\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk> | |
| 704 | \<Longrightarrow> DD (\<lambda>w. f w - g w) z = DD f z - DD g z" | |
| 705 | unfolding complex_differentiable_def | |
| 706 | by (rule DERIV_imp_DD) (metis (poly_guards_query) DERIV_diff DERIV_imp_DD) | |
| 707 | ||
| 708 | lemma complex_derivative_mult: | |
| 709 | "\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk> | |
| 710 | \<Longrightarrow> DD (\<lambda>w. f w * g w) z = f z * DD g z + DD f z * g z" | |
| 711 | unfolding complex_differentiable_def | |
| 712 | by (rule DERIV_imp_DD) (metis DERIV_imp_DD DERIV_mult') | |
| 713 | ||
| 714 | lemma complex_derivative_cmult: | |
| 715 | "f complex_differentiable at z \<Longrightarrow> DD (\<lambda>w. c * f w) z = c * DD f z" | |
| 716 | unfolding complex_differentiable_def | |
| 717 | by (metis DERIV_cmult DERIV_imp_DD) | |
| 718 | ||
| 719 | lemma complex_derivative_cmult_right: | |
| 720 | "f complex_differentiable at z \<Longrightarrow> DD (\<lambda>w. f w * c) z = DD f z * c" | |
| 721 | unfolding complex_differentiable_def | |
| 722 | by (metis DERIV_cmult_right DERIV_imp_DD) | |
| 723 | ||
| 724 | lemma complex_derivative_transform_within_open: | |
| 725 | "\<lbrakk>f holomorphic_on s; g holomorphic_on s; open s; z \<in> s; \<And>w. w \<in> s \<Longrightarrow> f w = g w\<rbrakk> | |
| 726 | \<Longrightarrow> DD f z = DD g z" | |
| 727 | unfolding holomorphic_on_def | |
| 728 | by (rule DERIV_imp_DD) (metis has_derivative_within_open DERIV_imp_DD DERIV_transform_within_open) | |
| 729 | ||
| 730 | lemma complex_derivative_compose_linear: | |
| 731 | "f complex_differentiable at (c * z) \<Longrightarrow> DD (\<lambda>w. f (c * w)) z = c * DD f (c * z)" | |
| 732 | apply (rule DERIV_imp_DD) | |
| 733 | apply (simp add: DD_iff_derivative_complex_differentiable [symmetric]) | |
| 734 | apply (metis DERIV_chain' DERIV_cmult_Id comm_semiring_1_class.normalizing_semiring_rules(7)) | |
| 735 | done | |
| 736 | ||
| 737 | subsection{*Caratheodory characterization.*}
 | |
| 738 | ||
| 739 | (*REPLACE the original version. BUT IN WHICH FILE??*) | |
| 740 | thm Deriv.CARAT_DERIV | |
| 741 | ||
| 742 | lemma complex_differentiable_caratheodory_at: | |
| 743 | "f complex_differentiable (at z) \<longleftrightarrow> | |
| 744 | (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z) g)" | |
| 745 | using CARAT_DERIV [of f] | |
| 746 | by (simp add: complex_differentiable_def has_field_derivative_def) | |
| 747 | ||
| 748 | lemma complex_differentiable_caratheodory_within: | |
| 749 | "f complex_differentiable (at z within s) \<longleftrightarrow> | |
| 750 | (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z within s) g)" | |
| 751 | using DERIV_caratheodory_within [of f] | |
| 752 | by (simp add: complex_differentiable_def has_field_derivative_def) | |
| 753 | ||
| 754 | subsection{*analyticity on a set*}
 | |
| 755 | ||
| 756 | definition analytic_on (infixl "(analytic'_on)" 50) | |
| 757 | where | |
| 758 | "f analytic_on s \<equiv> \<forall>x \<in> s. \<exists>e. 0 < e \<and> f holomorphic_on (ball x e)" | |
| 759 | ||
| 760 | lemma analytic_imp_holomorphic: | |
| 761 | "f analytic_on s \<Longrightarrow> f holomorphic_on s" | |
| 762 | unfolding analytic_on_def holomorphic_on_def | |
| 763 | apply (simp add: has_derivative_within_open [OF _ open_ball]) | |
| 764 | apply (metis DERIV_subset dist_self mem_ball top_greatest) | |
| 765 | done | |
| 766 | ||
| 767 | lemma analytic_on_open: | |
| 768 | "open s \<Longrightarrow> f analytic_on s \<longleftrightarrow> f holomorphic_on s" | |
| 769 | apply (auto simp: analytic_imp_holomorphic) | |
| 770 | apply (auto simp: analytic_on_def holomorphic_on_def) | |
| 771 | by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball) | |
| 772 | ||
| 773 | lemma analytic_on_imp_differentiable_at: | |
| 774 | "f analytic_on s \<Longrightarrow> x \<in> s \<Longrightarrow> f complex_differentiable (at x)" | |
| 775 | apply (auto simp: analytic_on_def holomorphic_on_differentiable) | |
| 776 | by (metis Topology_Euclidean_Space.open_ball centre_in_ball complex_differentiable_within_open) | |
| 777 | ||
| 778 | lemma analytic_on_subset: | |
| 779 | "f analytic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f analytic_on t" | |
| 780 | by (auto simp: analytic_on_def) | |
| 781 | ||
| 782 | lemma analytic_on_Un: | |
| 783 | "f analytic_on (s \<union> t) \<longleftrightarrow> f analytic_on s \<and> f analytic_on t" | |
| 784 | by (auto simp: analytic_on_def) | |
| 785 | ||
| 786 | lemma analytic_on_Union: | |
| 787 | "f analytic_on (\<Union> s) \<longleftrightarrow> (\<forall>t \<in> s. f analytic_on t)" | |
| 788 | by (auto simp: analytic_on_def) | |
| 789 | ||
| 790 | lemma analytic_on_holomorphic: | |
| 791 | "f analytic_on s \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f holomorphic_on t)" | |
| 792 | (is "?lhs = ?rhs") | |
| 793 | proof - | |
| 794 | have "?lhs \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t)" | |
| 795 | proof safe | |
| 796 | assume "f analytic_on s" | |
| 797 | then show "\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t" | |
| 798 | apply (simp add: analytic_on_def) | |
| 799 |       apply (rule exI [where x="\<Union>{u. open u \<and> f analytic_on u}"], auto)
 | |
| 800 | apply (metis Topology_Euclidean_Space.open_ball analytic_on_open centre_in_ball) | |
| 801 | by (metis analytic_on_def) | |
| 802 | next | |
| 803 | fix t | |
| 804 | assume "open t" "s \<subseteq> t" "f analytic_on t" | |
| 805 | then show "f analytic_on s" | |
| 806 | by (metis analytic_on_subset) | |
| 807 | qed | |
| 808 | also have "... \<longleftrightarrow> ?rhs" | |
| 809 | by (auto simp: analytic_on_open) | |
| 810 | finally show ?thesis . | |
| 811 | qed | |
| 812 | ||
| 813 | lemma analytic_on_linear: "(op * c) analytic_on s" | |
| 814 | apply (simp add: analytic_on_holomorphic holomorphic_on_linear) | |
| 815 | by (metis open_UNIV top_greatest) | |
| 816 | ||
| 817 | lemma analytic_on_const: "(\<lambda>z. c) analytic_on s" | |
| 818 | unfolding analytic_on_def | |
| 819 | by (metis holomorphic_on_const zero_less_one) | |
| 820 | ||
| 821 | lemma analytic_on_id: "id analytic_on s" | |
| 822 | unfolding analytic_on_def | |
| 823 | apply (simp add: holomorphic_on_id) | |
| 824 | by (metis gt_ex) | |
| 825 | ||
| 826 | lemma analytic_on_compose: | |
| 827 | assumes f: "f analytic_on s" | |
| 828 | and g: "g analytic_on (f ` s)" | |
| 829 | shows "(g o f) analytic_on s" | |
| 830 | unfolding analytic_on_def | |
| 831 | proof (intro ballI) | |
| 832 | fix x | |
| 833 | assume x: "x \<in> s" | |
| 834 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f | |
| 835 | by (metis analytic_on_def) | |
| 836 | obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g | |
| 837 | by (metis analytic_on_def g image_eqI x) | |
| 838 | have "isCont f x" | |
| 839 | by (metis analytic_on_imp_differentiable_at complex_differentiable_imp_continuous_at f x) | |
| 840 | with e' obtain d where d: "0 < d" and fd: "f ` ball x d \<subseteq> ball (f x) e'" | |
| 841 | by (auto simp: continuous_at_ball) | |
| 842 | have "g \<circ> f holomorphic_on ball x (min d e)" | |
| 843 | apply (rule holomorphic_on_compose) | |
| 844 | apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 845 | by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball) | |
| 846 | then show "\<exists>e>0. g \<circ> f holomorphic_on ball x e" | |
| 847 | by (metis d e min_less_iff_conj) | |
| 848 | qed | |
| 849 | ||
| 850 | lemma analytic_on_compose_gen: | |
| 851 | "f analytic_on s \<Longrightarrow> g analytic_on t \<Longrightarrow> (\<And>z. z \<in> s \<Longrightarrow> f z \<in> t) | |
| 852 | \<Longrightarrow> g o f analytic_on s" | |
| 853 | by (metis analytic_on_compose analytic_on_subset image_subset_iff) | |
| 854 | ||
| 855 | lemma analytic_on_neg: | |
| 856 | "f analytic_on s \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on s" | |
| 857 | by (metis analytic_on_holomorphic holomorphic_on_minus) | |
| 858 | ||
| 859 | lemma analytic_on_add: | |
| 860 | assumes f: "f analytic_on s" | |
| 861 | and g: "g analytic_on s" | |
| 862 | shows "(\<lambda>z. f z + g z) analytic_on s" | |
| 863 | unfolding analytic_on_def | |
| 864 | proof (intro ballI) | |
| 865 | fix z | |
| 866 | assume z: "z \<in> s" | |
| 867 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f | |
| 868 | by (metis analytic_on_def) | |
| 869 | obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g | |
| 870 | by (metis analytic_on_def g z) | |
| 871 | have "(\<lambda>z. f z + g z) holomorphic_on ball z (min e e')" | |
| 872 | apply (rule holomorphic_on_add) | |
| 873 | apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 874 | by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 875 | then show "\<exists>e>0. (\<lambda>z. f z + g z) holomorphic_on ball z e" | |
| 876 | by (metis e e' min_less_iff_conj) | |
| 877 | qed | |
| 878 | ||
| 879 | lemma analytic_on_diff: | |
| 880 | assumes f: "f analytic_on s" | |
| 881 | and g: "g analytic_on s" | |
| 882 | shows "(\<lambda>z. f z - g z) analytic_on s" | |
| 883 | unfolding analytic_on_def | |
| 884 | proof (intro ballI) | |
| 885 | fix z | |
| 886 | assume z: "z \<in> s" | |
| 887 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f | |
| 888 | by (metis analytic_on_def) | |
| 889 | obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g | |
| 890 | by (metis analytic_on_def g z) | |
| 891 | have "(\<lambda>z. f z - g z) holomorphic_on ball z (min e e')" | |
| 892 | apply (rule holomorphic_on_diff) | |
| 893 | apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 894 | by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 895 | then show "\<exists>e>0. (\<lambda>z. f z - g z) holomorphic_on ball z e" | |
| 896 | by (metis e e' min_less_iff_conj) | |
| 897 | qed | |
| 898 | ||
| 899 | lemma analytic_on_mult: | |
| 900 | assumes f: "f analytic_on s" | |
| 901 | and g: "g analytic_on s" | |
| 902 | shows "(\<lambda>z. f z * g z) analytic_on s" | |
| 903 | unfolding analytic_on_def | |
| 904 | proof (intro ballI) | |
| 905 | fix z | |
| 906 | assume z: "z \<in> s" | |
| 907 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f | |
| 908 | by (metis analytic_on_def) | |
| 909 | obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g | |
| 910 | by (metis analytic_on_def g z) | |
| 911 | have "(\<lambda>z. f z * g z) holomorphic_on ball z (min e e')" | |
| 912 | apply (rule holomorphic_on_mult) | |
| 913 | apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 914 | by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 915 | then show "\<exists>e>0. (\<lambda>z. f z * g z) holomorphic_on ball z e" | |
| 916 | by (metis e e' min_less_iff_conj) | |
| 917 | qed | |
| 918 | ||
| 919 | lemma analytic_on_inverse: | |
| 920 | assumes f: "f analytic_on s" | |
| 921 | and nz: "(\<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0)" | |
| 922 | shows "(\<lambda>z. inverse (f z)) analytic_on s" | |
| 923 | unfolding analytic_on_def | |
| 924 | proof (intro ballI) | |
| 925 | fix z | |
| 926 | assume z: "z \<in> s" | |
| 927 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f | |
| 928 | by (metis analytic_on_def) | |
| 929 | have "continuous_on (ball z e) f" | |
| 930 | by (metis fh holomorphic_on_imp_continuous_on) | |
| 931 | then obtain e' where e': "0 < e'" and nz': "\<And>y. dist z y < e' \<Longrightarrow> f y \<noteq> 0" | |
| 932 | by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz) | |
| 933 | have "(\<lambda>z. inverse (f z)) holomorphic_on ball z (min e e')" | |
| 934 | apply (rule holomorphic_on_inverse) | |
| 935 | apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball) | |
| 936 | by (metis nz' mem_ball min_less_iff_conj) | |
| 937 | then show "\<exists>e>0. (\<lambda>z. inverse (f z)) holomorphic_on ball z e" | |
| 938 | by (metis e e' min_less_iff_conj) | |
| 939 | qed | |
| 940 | ||
| 941 | ||
| 942 | lemma analytic_on_divide: | |
| 943 | assumes f: "f analytic_on s" | |
| 944 | and g: "g analytic_on s" | |
| 945 | and nz: "(\<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0)" | |
| 946 | shows "(\<lambda>z. f z / g z) analytic_on s" | |
| 947 | unfolding divide_inverse | |
| 948 | by (metis analytic_on_inverse analytic_on_mult f g nz) | |
| 949 | ||
| 950 | lemma analytic_on_power: | |
| 951 | "f analytic_on s \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on s" | |
| 952 | by (induct n) (auto simp: analytic_on_const analytic_on_mult) | |
| 953 | ||
| 954 | lemma analytic_on_setsum: | |
| 955 | "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s) | |
| 956 | \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) analytic_on s" | |
| 957 | by (induct I rule: finite_induct) (auto simp: analytic_on_const analytic_on_add) | |
| 958 | ||
| 959 | subsection{*analyticity at a point.*}
 | |
| 960 | ||
| 961 | lemma analytic_at_ball: | |
| 962 |   "f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)"
 | |
| 963 | by (metis analytic_on_def singleton_iff) | |
| 964 | ||
| 965 | lemma analytic_at: | |
| 966 |     "f analytic_on {z} \<longleftrightarrow> (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s)"
 | |
| 967 | by (metis analytic_on_holomorphic empty_subsetI insert_subset) | |
| 968 | ||
| 969 | lemma analytic_on_analytic_at: | |
| 970 |     "f analytic_on s \<longleftrightarrow> (\<forall>z \<in> s. f analytic_on {z})"
 | |
| 971 | by (metis analytic_at_ball analytic_on_def) | |
| 972 | ||
| 973 | lemma analytic_at_two: | |
| 974 |   "f analytic_on {z} \<and> g analytic_on {z} \<longleftrightarrow>
 | |
| 975 | (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s \<and> g holomorphic_on s)" | |
| 976 | (is "?lhs = ?rhs") | |
| 977 | proof | |
| 978 | assume ?lhs | |
| 979 | then obtain s t | |
| 980 | where st: "open s" "z \<in> s" "f holomorphic_on s" | |
| 981 | "open t" "z \<in> t" "g holomorphic_on t" | |
| 982 | by (auto simp: analytic_at) | |
| 983 | show ?rhs | |
| 984 | apply (rule_tac x="s \<inter> t" in exI) | |
| 985 | using st | |
| 986 | apply (auto simp: Diff_subset holomorphic_on_subset) | |
| 987 | done | |
| 988 | next | |
| 989 | assume ?rhs | |
| 990 | then show ?lhs | |
| 991 | by (force simp add: analytic_at) | |
| 992 | qed | |
| 993 | ||
| 994 | subsection{*Combining theorems for derivative with ``analytic at'' hypotheses*}
 | |
| 995 | ||
| 996 | lemma | |
| 997 |   assumes "f analytic_on {z}" "g analytic_on {z}"
 | |
| 998 | shows complex_derivative_add_at: "DD (\<lambda>w. f w + g w) z = DD f z + DD g z" | |
| 999 | and complex_derivative_diff_at: "DD (\<lambda>w. f w - g w) z = DD f z - DD g z" | |
| 1000 | and complex_derivative_mult_at: "DD (\<lambda>w. f w * g w) z = | |
| 1001 | f z * DD g z + DD f z * g z" | |
| 1002 | proof - | |
| 1003 | obtain s where s: "open s" "z \<in> s" "f holomorphic_on s" "g holomorphic_on s" | |
| 1004 | using assms by (metis analytic_at_two) | |
| 1005 | show "DD (\<lambda>w. f w + g w) z = DD f z + DD g z" | |
| 1006 | apply (rule DERIV_imp_DD [OF DERIV_add]) | |
| 1007 | using s | |
| 1008 | apply (auto simp: holomorphic_on_open complex_differentiable_def DD_iff_derivative_complex_differentiable) | |
| 1009 | done | |
| 1010 | show "DD (\<lambda>w. f w - g w) z = DD f z - DD g z" | |
| 1011 | apply (rule DERIV_imp_DD [OF DERIV_diff]) | |
| 1012 | using s | |
| 1013 | apply (auto simp: holomorphic_on_open complex_differentiable_def DD_iff_derivative_complex_differentiable) | |
| 1014 | done | |
| 1015 | show "DD (\<lambda>w. f w * g w) z = f z * DD g z + DD f z * g z" | |
| 1016 | apply (rule DERIV_imp_DD [OF DERIV_mult']) | |
| 1017 | using s | |
| 1018 | apply (auto simp: holomorphic_on_open complex_differentiable_def DD_iff_derivative_complex_differentiable) | |
| 1019 | done | |
| 1020 | qed | |
| 1021 | ||
| 1022 | lemma complex_derivative_cmult_at: | |
| 1023 |   "f analytic_on {z} \<Longrightarrow>  DD (\<lambda>w. c * f w) z = c * DD f z"
 | |
| 1024 | by (auto simp: complex_derivative_mult_at complex_derivative_const analytic_on_const) | |
| 1025 | ||
| 1026 | lemma complex_derivative_cmult_right_at: | |
| 1027 |   "f analytic_on {z} \<Longrightarrow>  DD (\<lambda>w. f w * c) z = DD f z * c"
 | |
| 1028 | by (auto simp: complex_derivative_mult_at complex_derivative_const analytic_on_const) | |
| 1029 | ||
| 1030 | text{*A composition lemma for functions of mixed type*}
 | |
| 1031 | lemma has_vector_derivative_real_complex: | |
| 1032 | fixes f :: "complex \<Rightarrow> complex" | |
| 1033 | assumes "DERIV f (of_real a) :> f'" | |
| 1034 | shows "((\<lambda>x. f (of_real x)) has_vector_derivative f') (at a)" | |
| 1035 | using diff_chain_at [OF has_derivative_ident, of f "op * f'" a] assms | |
| 1036 | unfolding has_field_derivative_def has_vector_derivative_def o_def | |
| 1037 | by (auto simp: mult_ac scaleR_conv_of_real) | |
| 1038 | ||
| 1039 | subsection{*Complex differentiation of sequences and series*}
 | |
| 1040 | ||
| 1041 | lemma has_complex_derivative_sequence: | |
| 1042 | fixes s :: "complex set" | |
| 1043 | assumes cvs: "convex s" | |
| 1044 | and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" | |
| 1045 | and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s \<longrightarrow> norm (f' n x - g' x) \<le> e" | |
| 1046 | and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) ---> l) sequentially" | |
| 1047 | shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) ---> g x) sequentially \<and> | |
| 1048 | (g has_field_derivative (g' x)) (at x within s)" | |
| 1049 | proof - | |
| 1050 | from assms obtain x l where x: "x \<in> s" and tf: "((\<lambda>n. f n x) ---> l) sequentially" | |
| 1051 | by blast | |
| 1052 |   { fix e::real assume e: "e > 0"
 | |
| 1053 | then obtain N where N: "\<forall>n\<ge>N. \<forall>x. x \<in> s \<longrightarrow> cmod (f' n x - g' x) \<le> e" | |
| 1054 | by (metis conv) | |
| 1055 | have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h" | |
| 1056 | proof (rule exI [of _ N], clarify) | |
| 1057 | fix n y h | |
| 1058 | assume "N \<le> n" "y \<in> s" | |
| 1059 | then have "cmod (f' n y - g' y) \<le> e" | |
| 1060 | by (metis N) | |
| 1061 | then have "cmod h * cmod (f' n y - g' y) \<le> cmod h * e" | |
| 1062 | by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) | |
| 1063 | then show "cmod (f' n y * h - g' y * h) \<le> e * cmod h" | |
| 1064 | by (simp add: norm_mult [symmetric] field_simps) | |
| 1065 | qed | |
| 1066 | } note ** = this | |
| 1067 | show ?thesis | |
| 1068 | unfolding has_field_derivative_def | |
| 1069 | proof (rule has_derivative_sequence [OF cvs _ _ x]) | |
| 1070 | show "\<forall>n. \<forall>x\<in>s. (f n has_derivative (op * (f' n x))) (at x within s)" | |
| 1071 | by (metis has_field_derivative_def df) | |
| 1072 | next show "(\<lambda>n. f n x) ----> l" | |
| 1073 | by (rule tf) | |
| 1074 | next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h" | |
| 1075 | by (blast intro: **) | |
| 1076 | qed | |
| 1077 | qed | |
| 1078 | ||
| 1079 | ||
| 1080 | lemma has_complex_derivative_series: | |
| 1081 | fixes s :: "complex set" | |
| 1082 | assumes cvs: "convex s" | |
| 1083 | and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" | |
| 1084 | and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s | |
| 1085 | \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e" | |
| 1086 | and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) sums l)" | |
| 1087 | shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((g has_field_derivative g' x) (at x within s))" | |
| 1088 | proof - | |
| 1089 | from assms obtain x l where x: "x \<in> s" and sf: "((\<lambda>n. f n x) sums l)" | |
| 1090 | by blast | |
| 1091 |   { fix e::real assume e: "e > 0"
 | |
| 1092 | then obtain N where N: "\<forall>n x. n \<ge> N \<longrightarrow> x \<in> s | |
| 1093 | \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e" | |
| 1094 | by (metis conv) | |
| 1095 | have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h" | |
| 1096 | proof (rule exI [of _ N], clarify) | |
| 1097 | fix n y h | |
| 1098 | assume "N \<le> n" "y \<in> s" | |
| 1099 | then have "cmod ((\<Sum>i<n. f' i y) - g' y) \<le> e" | |
| 1100 | by (metis N) | |
| 1101 | then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e" | |
| 1102 | by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) | |
| 1103 | then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h" | |
| 1104 | by (simp add: norm_mult [symmetric] field_simps setsum_right_distrib) | |
| 1105 | qed | |
| 1106 | } note ** = this | |
| 1107 | show ?thesis | |
| 1108 | unfolding has_field_derivative_def | |
| 1109 | proof (rule has_derivative_series [OF cvs _ _ x]) | |
| 1110 | fix n x | |
| 1111 | assume "x \<in> s" | |
| 1112 | then show "((f n) has_derivative (\<lambda>z. z * f' n x)) (at x within s)" | |
| 1113 | by (metis df has_field_derivative_def mult_commute_abs) | |
| 1114 | next show " ((\<lambda>n. f n x) sums l)" | |
| 1115 | by (rule sf) | |
| 1116 | next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h" | |
| 1117 | by (blast intro: **) | |
| 1118 | qed | |
| 1119 | qed | |
| 1120 | ||
| 1121 | subsection{*Bound theorem*}
 | |
| 1122 | ||
| 1123 | lemma complex_differentiable_bound: | |
| 1124 | fixes s :: "complex set" | |
| 1125 | assumes cvs: "convex s" | |
| 1126 | and df: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z within s)" | |
| 1127 | and dn: "\<And>z. z \<in> s \<Longrightarrow> norm (f' z) \<le> B" | |
| 1128 | and "x \<in> s" "y \<in> s" | |
| 1129 | shows "norm(f x - f y) \<le> B * norm(x - y)" | |
| 1130 | apply (rule differentiable_bound [OF cvs]) | |
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 1131 | apply (rule ballI, erule df [unfolded has_field_derivative_def]) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 1132 | apply (rule ballI, rule onorm_le, simp add: norm_mult mult_right_mono dn) | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 1133 | apply fact | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 1134 | apply fact | 
| 56215 | 1135 | done | 
| 1136 | ||
| 1137 | subsection{*Inverse function theorem for complex derivatives.*}
 | |
| 1138 | ||
| 1139 | lemma has_complex_derivative_inverse_basic: | |
| 1140 | fixes f :: "complex \<Rightarrow> complex" | |
| 1141 | shows "DERIV f (g y) :> f' \<Longrightarrow> | |
| 1142 | f' \<noteq> 0 \<Longrightarrow> | |
| 1143 | continuous (at y) g \<Longrightarrow> | |
| 1144 | open t \<Longrightarrow> | |
| 1145 | y \<in> t \<Longrightarrow> | |
| 1146 | (\<And>z. z \<in> t \<Longrightarrow> f (g z) = z) | |
| 1147 | \<Longrightarrow> DERIV g y :> inverse (f')" | |
| 1148 | unfolding has_field_derivative_def | |
| 1149 | apply (rule has_derivative_inverse_basic) | |
| 1150 | apply (auto simp: bounded_linear_mult_right) | |
| 1151 | done | |
| 1152 | ||
| 1153 | (*Used only once, in Multivariate/cauchy.ml. *) | |
| 1154 | lemma has_complex_derivative_inverse_strong: | |
| 1155 | fixes f :: "complex \<Rightarrow> complex" | |
| 1156 | shows "DERIV f x :> f' \<Longrightarrow> | |
| 1157 | f' \<noteq> 0 \<Longrightarrow> | |
| 1158 | open s \<Longrightarrow> | |
| 1159 | x \<in> s \<Longrightarrow> | |
| 1160 | continuous_on s f \<Longrightarrow> | |
| 1161 | (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z) | |
| 1162 | \<Longrightarrow> DERIV g (f x) :> inverse (f')" | |
| 1163 | unfolding has_field_derivative_def | |
| 1164 | apply (rule has_derivative_inverse_strong [of s x f g ]) | |
| 1165 | using assms | |
| 1166 | by auto | |
| 1167 | ||
| 1168 | lemma has_complex_derivative_inverse_strong_x: | |
| 1169 | fixes f :: "complex \<Rightarrow> complex" | |
| 1170 | shows "DERIV f (g y) :> f' \<Longrightarrow> | |
| 1171 | f' \<noteq> 0 \<Longrightarrow> | |
| 1172 | open s \<Longrightarrow> | |
| 1173 | continuous_on s f \<Longrightarrow> | |
| 1174 | g y \<in> s \<Longrightarrow> f(g y) = y \<Longrightarrow> | |
| 1175 | (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z) | |
| 1176 | \<Longrightarrow> DERIV g y :> inverse (f')" | |
| 1177 | unfolding has_field_derivative_def | |
| 1178 | apply (rule has_derivative_inverse_strong_x [of s g y f]) | |
| 1179 | using assms | |
| 1180 | by auto | |
| 1181 | ||
| 1182 | subsection{*Further useful properties of complex conjugation*}
 | |
| 1183 | ||
| 1184 | lemma lim_cnj: "((\<lambda>x. cnj(f x)) ---> cnj l) F \<longleftrightarrow> (f ---> l) F" | |
| 1185 | by (simp add: tendsto_iff dist_complex_def Complex.complex_cnj_diff [symmetric]) | |
| 1186 | ||
| 1187 | lemma sums_cnj: "((\<lambda>x. cnj(f x)) sums cnj l) \<longleftrightarrow> (f sums l)" | |
| 1188 | by (simp add: sums_def lim_cnj cnj_setsum [symmetric]) | |
| 1189 | ||
| 1190 | lemma continuous_within_cnj: "continuous (at z within s) cnj" | |
| 1191 | by (metis bounded_linear_cnj linear_continuous_within) | |
| 1192 | ||
| 1193 | lemma continuous_on_cnj: "continuous_on s cnj" | |
| 1194 | by (metis bounded_linear_cnj linear_continuous_on) | |
| 1195 | ||
| 1196 | subsection{*Some limit theorems about real part of real series etc.*}
 | |
| 1197 | ||
| 1198 | lemma real_lim: | |
| 1199 | fixes l::complex | |
| 1200 | assumes "(f ---> l) F" and " ~(trivial_limit F)" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>" | |
| 1201 | shows "l \<in> \<real>" | |
| 1202 | proof - | |
| 1203 | have 1: "((\<lambda>i. Im (f i)) ---> Im l) F" | |
| 1204 | by (metis assms(1) tendsto_Im) | |
| 56217 
dc429a5b13c4
Some rationalisation of basic lemmas
 paulson <lp15@cam.ac.uk> parents: 
56215diff
changeset | 1205 | then have "((\<lambda>i. Im (f i)) ---> 0) F" using assms | 
| 
dc429a5b13c4
Some rationalisation of basic lemmas
 paulson <lp15@cam.ac.uk> parents: 
56215diff
changeset | 1206 | by (metis (mono_tags, lifting) Lim_eventually complex_is_Real_iff eventually_mono) | 
| 56215 | 1207 | then show ?thesis using 1 | 
| 1208 | by (metis 1 assms(2) complex_is_Real_iff tendsto_unique) | |
| 1209 | qed | |
| 1210 | ||
| 1211 | lemma real_lim_sequentially: | |
| 1212 | fixes l::complex | |
| 1213 | shows "(f ---> l) sequentially \<Longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>" | |
| 1214 | by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially) | |
| 1215 | ||
| 1216 | lemma real_series: | |
| 1217 | fixes l::complex | |
| 1218 | shows "f sums l \<Longrightarrow> (\<And>n. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>" | |
| 1219 | unfolding sums_def | |
| 1220 | by (metis real_lim_sequentially setsum_in_Reals) | |
| 1221 | ||
| 1222 | ||
| 1223 | lemma Lim_null_comparison_Re: | |
| 1224 | "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F \<Longrightarrow> (g ---> 0) F \<Longrightarrow> (f ---> 0) F" | |
| 1225 | by (metis Lim_null_comparison complex_Re_zero tendsto_Re) | |
| 1226 | ||
| 1227 | ||
| 1228 | lemma norm_setsum_bound: | |
| 1229 | fixes n::nat | |
| 1230 | shows" \<lbrakk>\<And>n. N \<le> n \<Longrightarrow> norm (f n) \<le> g n; N \<le> m\<rbrakk> | |
| 1231 |        \<Longrightarrow> norm (setsum f {m..<n}) \<le> setsum g {m..<n}"
 | |
| 1232 | apply (induct n, auto) | |
| 1233 | by (metis dual_order.trans le_cases le_neq_implies_less norm_triangle_mono) | |
| 1234 | ||
| 1235 | ||
| 56217 
dc429a5b13c4
Some rationalisation of basic lemmas
 paulson <lp15@cam.ac.uk> parents: 
56215diff
changeset | 1236 | (*MOVE? But not to Finite_Cartesian_Product*) | 
| 56215 | 1237 | lemma sums_vec_nth : | 
| 1238 | assumes "f sums a" | |
| 1239 | shows "(\<lambda>x. f x $ i) sums a $ i" | |
| 1240 | using assms unfolding sums_def | |
| 1241 | by (auto dest: tendsto_vec_nth [where i=i]) | |
| 1242 | ||
| 1243 | lemma summable_vec_nth : | |
| 1244 | assumes "summable f" | |
| 1245 | shows "summable (\<lambda>x. f x $ i)" | |
| 1246 | using assms unfolding summable_def | |
| 1247 | by (blast intro: sums_vec_nth) | |
| 1248 | ||
| 1249 | lemma sums_Re: | |
| 1250 | assumes "f sums a" | |
| 1251 | shows "(\<lambda>x. Re (f x)) sums Re a" | |
| 1252 | using assms | |
| 1253 | by (auto simp: sums_def Re_setsum [symmetric] isCont_tendsto_compose [of a Re]) | |
| 1254 | ||
| 1255 | lemma sums_Im: | |
| 1256 | assumes "f sums a" | |
| 1257 | shows "(\<lambda>x. Im (f x)) sums Im a" | |
| 1258 | using assms | |
| 1259 | by (auto simp: sums_def Im_setsum [symmetric] isCont_tendsto_compose [of a Im]) | |
| 1260 | ||
| 1261 | lemma summable_Re: | |
| 1262 | assumes "summable f" | |
| 1263 | shows "summable (\<lambda>x. Re (f x))" | |
| 1264 | using assms unfolding summable_def | |
| 1265 | by (blast intro: sums_Re) | |
| 1266 | ||
| 1267 | lemma summable_Im: | |
| 1268 | assumes "summable f" | |
| 1269 | shows "summable (\<lambda>x. Im (f x))" | |
| 1270 | using assms unfolding summable_def | |
| 1271 | by (blast intro: sums_Im) | |
| 1272 | ||
| 1273 | lemma series_comparison_complex: | |
| 1274 | fixes f:: "nat \<Rightarrow> 'a::banach" | |
| 1275 | assumes sg: "summable g" | |
| 1276 | and "\<And>n. g n \<in> \<real>" "\<And>n. Re (g n) \<ge> 0" | |
| 1277 | and fg: "\<And>n. n \<ge> N \<Longrightarrow> norm(f n) \<le> norm(g n)" | |
| 1278 | shows "summable f" | |
| 1279 | proof - | |
| 1280 | have g: "\<And>n. cmod (g n) = Re (g n)" using assms | |
| 1281 | by (metis abs_of_nonneg in_Reals_norm) | |
| 1282 | show ?thesis | |
| 56217 
dc429a5b13c4
Some rationalisation of basic lemmas
 paulson <lp15@cam.ac.uk> parents: 
56215diff
changeset | 1283 | apply (rule summable_comparison_test' [where g = "\<lambda>n. norm (g n)" and N=N]) | 
| 56215 | 1284 | using sg | 
| 1285 | apply (auto simp: summable_def) | |
| 1286 | apply (rule_tac x="Re s" in exI) | |
| 1287 | apply (auto simp: g sums_Re) | |
| 1288 | apply (metis fg g) | |
| 1289 | done | |
| 1290 | qed | |
| 1291 | ||
| 1292 | lemma summable_complex_of_real [simp]: | |
| 1293 | "summable (\<lambda>n. complex_of_real (f n)) = summable f" | |
| 1294 | apply (auto simp: Series.summable_Cauchy) | |
| 1295 | apply (drule_tac x = e in spec, auto) | |
| 1296 | apply (rule_tac x=N in exI) | |
| 1297 | apply (auto simp: of_real_setsum [symmetric]) | |
| 1298 | done | |
| 1299 | ||
| 1300 | ||
| 1301 | ||
| 1302 | ||
| 1303 | lemma setsum_Suc_reindex: | |
| 1304 | fixes f :: "nat \<Rightarrow> 'a::ab_group_add" | |
| 1305 |     shows  "setsum f {0..n} = f 0 - f (Suc n) + setsum (\<lambda>i. f (Suc i)) {0..n}"
 | |
| 1306 | by (induct n) auto | |
| 1307 | ||
| 1308 | ||
| 56217 
dc429a5b13c4
Some rationalisation of basic lemmas
 paulson <lp15@cam.ac.uk> parents: 
56215diff
changeset | 1309 | text{*A kind of complex Taylor theorem.*}
 | 
| 56215 | 1310 | lemma complex_taylor: | 
| 1311 | assumes s: "convex s" | |
| 1312 | and f: "\<And>i x. x \<in> s \<Longrightarrow> i \<le> n \<Longrightarrow> (f i has_field_derivative f (Suc i) x) (at x within s)" | |
| 1313 | and B: "\<And>x. x \<in> s \<Longrightarrow> cmod (f (Suc n) x) \<le> B" | |
| 1314 | and w: "w \<in> s" | |
| 1315 | and z: "z \<in> s" | |
| 1316 | shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / of_nat (fact i))) | |
| 1317 | \<le> B * cmod(z - w)^(Suc n) / fact n" | |
| 1318 | proof - | |
| 1319 | have wzs: "closed_segment w z \<subseteq> s" using assms | |
| 1320 | by (metis convex_contains_segment) | |
| 1321 |   { fix u
 | |
| 1322 | assume "u \<in> closed_segment w z" | |
| 1323 | then have "u \<in> s" | |
| 1324 | by (metis wzs subsetD) | |
| 1325 | have "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / of_nat (fact i) + | |
| 1326 | f (Suc i) u * (z-u)^i / of_nat (fact i)) = | |
| 1327 | f (Suc n) u * (z-u) ^ n / of_nat (fact n)" | |
| 1328 | proof (induction n) | |
| 1329 | case 0 show ?case by simp | |
| 1330 | next | |
| 1331 | case (Suc n) | |
| 1332 | have "(\<Sum>i\<le>Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / of_nat (fact i) + | |
| 1333 | f (Suc i) u * (z-u) ^ i / of_nat (fact i)) = | |
| 1334 | f (Suc n) u * (z-u) ^ n / of_nat (fact n) + | |
| 1335 | f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / of_nat (fact (Suc n)) - | |
| 1336 | f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / of_nat (fact (Suc n))" | |
| 1337 | using Suc by simp | |
| 1338 | also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / of_nat (fact (Suc n))" | |
| 1339 | proof - | |
| 1340 | have "of_nat(fact(Suc n)) * | |
| 1341 | (f(Suc n) u *(z-u) ^ n / of_nat(fact n) + | |
| 1342 | f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / of_nat(fact(Suc n)) - | |
| 1343 | f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / of_nat(fact(Suc n))) = | |
| 1344 | (of_nat(fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / of_nat(fact n) + | |
| 1345 | (of_nat(fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / of_nat(fact(Suc n))) - | |
| 1346 | (of_nat(fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / of_nat(fact(Suc n))" | |
| 1347 | by (simp add: algebra_simps del: fact_Suc) | |
| 1348 | also have "... = | |
| 1349 | (of_nat (fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / of_nat (fact n) + | |
| 1350 | (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - | |
| 1351 | (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" | |
| 1352 | by (simp del: fact_Suc) | |
| 1353 | also have "... = | |
| 1354 | (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) + | |
| 1355 | (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - | |
| 1356 | (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" | |
| 1357 | by (simp only: fact_Suc of_nat_mult mult_ac) simp | |
| 1358 | also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)" | |
| 1359 | by (simp add: algebra_simps) | |
| 1360 | finally show ?thesis | |
| 1361 | by (simp add: mult_left_cancel [where c = "of_nat (fact (Suc n))", THEN iffD1] del: fact_Suc) | |
| 1362 | qed | |
| 1363 | finally show ?case . | |
| 1364 | qed | |
| 1365 | then have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / of_nat (fact i))) | |
| 1366 | has_field_derivative f (Suc n) u * (z-u) ^ n / of_nat (fact n)) | |
| 1367 | (at u within s)" | |
| 1368 | apply (intro DERIV_intros DERIV_power[THEN DERIV_cong]) | |
| 1369 | apply (blast intro: assms `u \<in> s`) | |
| 1370 | apply (rule refl)+ | |
| 1371 | apply (auto simp: field_simps) | |
| 1372 | done | |
| 1373 | } note sum_deriv = this | |
| 1374 |   { fix u
 | |
| 1375 | assume u: "u \<in> closed_segment w z" | |
| 1376 | then have us: "u \<in> s" | |
| 1377 | by (metis wzs subsetD) | |
| 1378 | have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> cmod (f (Suc n) u) * cmod (u - z) ^ n" | |
| 1379 | by (metis norm_minus_commute order_refl) | |
| 1380 | also have "... \<le> cmod (f (Suc n) u) * cmod (z - w) ^ n" | |
| 1381 | by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u]) | |
| 1382 | also have "... \<le> B * cmod (z - w) ^ n" | |
| 1383 | by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us]) | |
| 1384 | finally have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> B * cmod (z - w) ^ n" . | |
| 1385 | } note cmod_bound = this | |
| 1386 | have "(\<Sum>i\<le>n. f i z * (z - z) ^ i / of_nat (fact i)) = (\<Sum>i\<le>n. (f i z / of_nat (fact i)) * 0 ^ i)" | |
| 1387 | by simp | |
| 1388 | also have "\<dots> = f 0 z / of_nat (fact 0)" | |
| 1389 | by (subst setsum_zero_power) simp | |
| 1390 | finally have "cmod (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / of_nat (fact i))) | |
| 1391 | \<le> cmod ((\<Sum>i\<le>n. f i w * (z - w) ^ i / of_nat (fact i)) - | |
| 1392 | (\<Sum>i\<le>n. f i z * (z - z) ^ i / of_nat (fact i)))" | |
| 1393 | by (simp add: norm_minus_commute) | |
| 1394 | also have "... \<le> B * cmod (z - w) ^ n / real_of_nat (fact n) * cmod (w - z)" | |
| 1395 | apply (rule complex_differentiable_bound | |
| 1396 | [where f' = "\<lambda>w. f (Suc n) w * (z - w)^n / of_nat(fact n)" | |
| 1397 | and s = "closed_segment w z", OF convex_segment]) | |
| 1398 | apply (auto simp: ends_in_segment real_of_nat_def DERIV_subset [OF sum_deriv wzs] | |
| 1399 | norm_divide norm_mult norm_power divide_le_cancel cmod_bound) | |
| 1400 | done | |
| 1401 | also have "... \<le> B * cmod (z - w) ^ Suc n / real (fact n)" | |
| 1402 | by (simp add: algebra_simps norm_minus_commute real_of_nat_def) | |
| 1403 | finally show ?thesis . | |
| 1404 | qed | |
| 1405 | ||
| 1406 | end |