|
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 section \<open>Complex Analysis Basics\<close> |
|
6 |
|
7 theory Complex_Analysis_Basics |
|
8 imports Cartesian_Euclidean_Space "~~/src/HOL/Library/Nonpos_Ints" |
|
9 begin |
|
10 |
|
11 |
|
12 subsection\<open>General lemmas\<close> |
|
13 |
|
14 lemma nonneg_Reals_cmod_eq_Re: "z \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> norm z = Re z" |
|
15 by (simp add: complex_nonneg_Reals_iff cmod_eq_Re) |
|
16 |
|
17 lemma has_derivative_mult_right: |
|
18 fixes c:: "'a :: real_normed_algebra" |
|
19 shows "((op * c) has_derivative (op * c)) F" |
|
20 by (rule has_derivative_mult_right [OF has_derivative_id]) |
|
21 |
|
22 lemma has_derivative_of_real[derivative_intros, simp]: |
|
23 "(f has_derivative f') F \<Longrightarrow> ((\<lambda>x. of_real (f x)) has_derivative (\<lambda>x. of_real (f' x))) F" |
|
24 using bounded_linear.has_derivative[OF bounded_linear_of_real] . |
|
25 |
|
26 lemma has_vector_derivative_real_complex: |
|
27 "DERIV f (of_real a) :> f' \<Longrightarrow> ((\<lambda>x. f (of_real x)) has_vector_derivative f') (at a within s)" |
|
28 using has_derivative_compose[of of_real of_real a _ f "op * f'"] |
|
29 by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) |
|
30 |
|
31 lemma fact_cancel: |
|
32 fixes c :: "'a::real_field" |
|
33 shows "of_nat (Suc n) * c / (fact (Suc n)) = c / (fact n)" |
|
34 by (simp add: of_nat_mult del: of_nat_Suc times_nat.simps) |
|
35 |
|
36 lemma bilinear_times: |
|
37 fixes c::"'a::real_algebra" shows "bilinear (\<lambda>x y::'a. x*y)" |
|
38 by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI) |
|
39 |
|
40 lemma linear_cnj: "linear cnj" |
|
41 using bounded_linear.linear[OF bounded_linear_cnj] . |
|
42 |
|
43 lemma tendsto_Re_upper: |
|
44 assumes "~ (trivial_limit F)" |
|
45 "(f \<longlongrightarrow> l) F" |
|
46 "eventually (\<lambda>x. Re(f x) \<le> b) F" |
|
47 shows "Re(l) \<le> b" |
|
48 by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Re) |
|
49 |
|
50 lemma tendsto_Re_lower: |
|
51 assumes "~ (trivial_limit F)" |
|
52 "(f \<longlongrightarrow> l) F" |
|
53 "eventually (\<lambda>x. b \<le> Re(f x)) F" |
|
54 shows "b \<le> Re(l)" |
|
55 by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Re) |
|
56 |
|
57 lemma tendsto_Im_upper: |
|
58 assumes "~ (trivial_limit F)" |
|
59 "(f \<longlongrightarrow> l) F" |
|
60 "eventually (\<lambda>x. Im(f x) \<le> b) F" |
|
61 shows "Im(l) \<le> b" |
|
62 by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Im) |
|
63 |
|
64 lemma tendsto_Im_lower: |
|
65 assumes "~ (trivial_limit F)" |
|
66 "(f \<longlongrightarrow> l) F" |
|
67 "eventually (\<lambda>x. b \<le> Im(f x)) F" |
|
68 shows "b \<le> Im(l)" |
|
69 by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Im) |
|
70 |
|
71 lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = op * 0" |
|
72 by auto |
|
73 |
|
74 lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = op * 1" |
|
75 by auto |
|
76 |
|
77 lemma continuous_mult_left: |
|
78 fixes c::"'a::real_normed_algebra" |
|
79 shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. c * f x)" |
|
80 by (rule continuous_mult [OF continuous_const]) |
|
81 |
|
82 lemma continuous_mult_right: |
|
83 fixes c::"'a::real_normed_algebra" |
|
84 shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x * c)" |
|
85 by (rule continuous_mult [OF _ continuous_const]) |
|
86 |
|
87 lemma continuous_on_mult_left: |
|
88 fixes c::"'a::real_normed_algebra" |
|
89 shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. c * f x)" |
|
90 by (rule continuous_on_mult [OF continuous_on_const]) |
|
91 |
|
92 lemma continuous_on_mult_right: |
|
93 fixes c::"'a::real_normed_algebra" |
|
94 shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x * c)" |
|
95 by (rule continuous_on_mult [OF _ continuous_on_const]) |
|
96 |
|
97 lemma uniformly_continuous_on_cmul_right [continuous_intros]: |
|
98 fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" |
|
99 shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. f x * c)" |
|
100 using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] . |
|
101 |
|
102 lemma uniformly_continuous_on_cmul_left[continuous_intros]: |
|
103 fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" |
|
104 assumes "uniformly_continuous_on s f" |
|
105 shows "uniformly_continuous_on s (\<lambda>x. c * f x)" |
|
106 by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right) |
|
107 |
|
108 lemma continuous_within_norm_id [continuous_intros]: "continuous (at x within S) norm" |
|
109 by (rule continuous_norm [OF continuous_ident]) |
|
110 |
|
111 lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm" |
|
112 by (intro continuous_on_id continuous_on_norm) |
|
113 |
|
114 subsection\<open>DERIV stuff\<close> |
|
115 |
|
116 lemma DERIV_zero_connected_constant: |
|
117 fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a" |
|
118 assumes "connected s" |
|
119 and "open s" |
|
120 and "finite k" |
|
121 and "continuous_on s f" |
|
122 and "\<forall>x\<in>(s - k). DERIV f x :> 0" |
|
123 obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c" |
|
124 using has_derivative_zero_connected_constant [OF assms(1-4)] assms |
|
125 by (metis DERIV_const has_derivative_const Diff_iff at_within_open frechet_derivative_at has_field_derivative_def) |
|
126 |
|
127 lemma DERIV_zero_constant: |
|
128 fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a" |
|
129 shows "\<lbrakk>convex s; |
|
130 \<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)\<rbrakk> |
|
131 \<Longrightarrow> \<exists>c. \<forall>x \<in> s. f(x) = c" |
|
132 by (auto simp: has_field_derivative_def lambda_zero intro: has_derivative_zero_constant) |
|
133 |
|
134 lemma DERIV_zero_unique: |
|
135 fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a" |
|
136 assumes "convex s" |
|
137 and d0: "\<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)" |
|
138 and "a \<in> s" |
|
139 and "x \<in> s" |
|
140 shows "f x = f a" |
|
141 by (rule has_derivative_zero_unique [OF assms(1) _ assms(4,3)]) |
|
142 (metis d0 has_field_derivative_imp_has_derivative lambda_zero) |
|
143 |
|
144 lemma DERIV_zero_connected_unique: |
|
145 fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a" |
|
146 assumes "connected s" |
|
147 and "open s" |
|
148 and d0: "\<And>x. x\<in>s \<Longrightarrow> DERIV f x :> 0" |
|
149 and "a \<in> s" |
|
150 and "x \<in> s" |
|
151 shows "f x = f a" |
|
152 by (rule has_derivative_zero_unique_connected [OF assms(2,1) _ assms(5,4)]) |
|
153 (metis has_field_derivative_def lambda_zero d0) |
|
154 |
|
155 lemma DERIV_transform_within: |
|
156 assumes "(f has_field_derivative f') (at a within s)" |
|
157 and "0 < d" "a \<in> s" |
|
158 and "\<And>x. x\<in>s \<Longrightarrow> dist x a < d \<Longrightarrow> f x = g x" |
|
159 shows "(g has_field_derivative f') (at a within s)" |
|
160 using assms unfolding has_field_derivative_def |
|
161 by (blast intro: has_derivative_transform_within) |
|
162 |
|
163 lemma DERIV_transform_within_open: |
|
164 assumes "DERIV f a :> f'" |
|
165 and "open s" "a \<in> s" |
|
166 and "\<And>x. x\<in>s \<Longrightarrow> f x = g x" |
|
167 shows "DERIV g a :> f'" |
|
168 using assms unfolding has_field_derivative_def |
|
169 by (metis has_derivative_transform_within_open) |
|
170 |
|
171 lemma DERIV_transform_at: |
|
172 assumes "DERIV f a :> f'" |
|
173 and "0 < d" |
|
174 and "\<And>x. dist x a < d \<Longrightarrow> f x = g x" |
|
175 shows "DERIV g a :> f'" |
|
176 by (blast intro: assms DERIV_transform_within) |
|
177 |
|
178 (*generalising DERIV_isconst_all, which requires type real (using the ordering)*) |
|
179 lemma DERIV_zero_UNIV_unique: |
|
180 fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a" |
|
181 shows "(\<And>x. DERIV f x :> 0) \<Longrightarrow> f x = f a" |
|
182 by (metis DERIV_zero_unique UNIV_I convex_UNIV) |
|
183 |
|
184 subsection \<open>Some limit theorems about real part of real series etc.\<close> |
|
185 |
|
186 (*MOVE? But not to Finite_Cartesian_Product*) |
|
187 lemma sums_vec_nth : |
|
188 assumes "f sums a" |
|
189 shows "(\<lambda>x. f x $ i) sums a $ i" |
|
190 using assms unfolding sums_def |
|
191 by (auto dest: tendsto_vec_nth [where i=i]) |
|
192 |
|
193 lemma summable_vec_nth : |
|
194 assumes "summable f" |
|
195 shows "summable (\<lambda>x. f x $ i)" |
|
196 using assms unfolding summable_def |
|
197 by (blast intro: sums_vec_nth) |
|
198 |
|
199 subsection \<open>Complex number lemmas\<close> |
|
200 |
|
201 lemma |
|
202 shows open_halfspace_Re_lt: "open {z. Re(z) < b}" |
|
203 and open_halfspace_Re_gt: "open {z. Re(z) > b}" |
|
204 and closed_halfspace_Re_ge: "closed {z. Re(z) \<ge> b}" |
|
205 and closed_halfspace_Re_le: "closed {z. Re(z) \<le> b}" |
|
206 and closed_halfspace_Re_eq: "closed {z. Re(z) = b}" |
|
207 and open_halfspace_Im_lt: "open {z. Im(z) < b}" |
|
208 and open_halfspace_Im_gt: "open {z. Im(z) > b}" |
|
209 and closed_halfspace_Im_ge: "closed {z. Im(z) \<ge> b}" |
|
210 and closed_halfspace_Im_le: "closed {z. Im(z) \<le> b}" |
|
211 and closed_halfspace_Im_eq: "closed {z. Im(z) = b}" |
|
212 by (intro open_Collect_less closed_Collect_le closed_Collect_eq continuous_on_Re |
|
213 continuous_on_Im continuous_on_id continuous_on_const)+ |
|
214 |
|
215 lemma closed_complex_Reals: "closed (\<real> :: complex set)" |
|
216 proof - |
|
217 have "(\<real> :: complex set) = {z. Im z = 0}" |
|
218 by (auto simp: complex_is_Real_iff) |
|
219 then show ?thesis |
|
220 by (metis closed_halfspace_Im_eq) |
|
221 qed |
|
222 |
|
223 lemma closed_Real_halfspace_Re_le: "closed (\<real> \<inter> {w. Re w \<le> x})" |
|
224 by (simp add: closed_Int closed_complex_Reals closed_halfspace_Re_le) |
|
225 |
|
226 corollary closed_nonpos_Reals_complex [simp]: "closed (\<real>\<^sub>\<le>\<^sub>0 :: complex set)" |
|
227 proof - |
|
228 have "\<real>\<^sub>\<le>\<^sub>0 = \<real> \<inter> {z. Re(z) \<le> 0}" |
|
229 using complex_nonpos_Reals_iff complex_is_Real_iff by auto |
|
230 then show ?thesis |
|
231 by (metis closed_Real_halfspace_Re_le) |
|
232 qed |
|
233 |
|
234 lemma closed_Real_halfspace_Re_ge: "closed (\<real> \<inter> {w. x \<le> Re(w)})" |
|
235 using closed_halfspace_Re_ge |
|
236 by (simp add: closed_Int closed_complex_Reals) |
|
237 |
|
238 corollary closed_nonneg_Reals_complex [simp]: "closed (\<real>\<^sub>\<ge>\<^sub>0 :: complex set)" |
|
239 proof - |
|
240 have "\<real>\<^sub>\<ge>\<^sub>0 = \<real> \<inter> {z. Re(z) \<ge> 0}" |
|
241 using complex_nonneg_Reals_iff complex_is_Real_iff by auto |
|
242 then show ?thesis |
|
243 by (metis closed_Real_halfspace_Re_ge) |
|
244 qed |
|
245 |
|
246 lemma closed_real_abs_le: "closed {w \<in> \<real>. \<bar>Re w\<bar> \<le> r}" |
|
247 proof - |
|
248 have "{w \<in> \<real>. \<bar>Re w\<bar> \<le> r} = (\<real> \<inter> {w. Re w \<le> r}) \<inter> (\<real> \<inter> {w. Re w \<ge> -r})" |
|
249 by auto |
|
250 then show "closed {w \<in> \<real>. \<bar>Re w\<bar> \<le> r}" |
|
251 by (simp add: closed_Int closed_Real_halfspace_Re_ge closed_Real_halfspace_Re_le) |
|
252 qed |
|
253 |
|
254 lemma real_lim: |
|
255 fixes l::complex |
|
256 assumes "(f \<longlongrightarrow> l) F" and "~(trivial_limit F)" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>" |
|
257 shows "l \<in> \<real>" |
|
258 proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)]) |
|
259 show "eventually (\<lambda>x. f x \<in> \<real>) F" |
|
260 using assms(3, 4) by (auto intro: eventually_mono) |
|
261 qed |
|
262 |
|
263 lemma real_lim_sequentially: |
|
264 fixes l::complex |
|
265 shows "(f \<longlongrightarrow> l) sequentially \<Longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>" |
|
266 by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially) |
|
267 |
|
268 lemma real_series: |
|
269 fixes l::complex |
|
270 shows "f sums l \<Longrightarrow> (\<And>n. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>" |
|
271 unfolding sums_def |
|
272 by (metis real_lim_sequentially setsum_in_Reals) |
|
273 |
|
274 lemma Lim_null_comparison_Re: |
|
275 assumes "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F" "(g \<longlongrightarrow> 0) F" shows "(f \<longlongrightarrow> 0) F" |
|
276 by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp |
|
277 |
|
278 subsection\<open>Holomorphic functions\<close> |
|
279 |
|
280 definition field_differentiable :: "['a \<Rightarrow> 'a::real_normed_field, 'a filter] \<Rightarrow> bool" |
|
281 (infixr "(field'_differentiable)" 50) |
|
282 where "f field_differentiable F \<equiv> \<exists>f'. (f has_field_derivative f') F" |
|
283 |
|
284 lemma field_differentiable_derivI: |
|
285 "f field_differentiable (at x) \<Longrightarrow> (f has_field_derivative deriv f x) (at x)" |
|
286 by (simp add: field_differentiable_def DERIV_deriv_iff_has_field_derivative) |
|
287 |
|
288 lemma field_differentiable_imp_continuous_at: |
|
289 "f field_differentiable (at x within s) \<Longrightarrow> continuous (at x within s) f" |
|
290 by (metis DERIV_continuous field_differentiable_def) |
|
291 |
|
292 lemma field_differentiable_within_subset: |
|
293 "\<lbrakk>f field_differentiable (at x within s); t \<subseteq> s\<rbrakk> |
|
294 \<Longrightarrow> f field_differentiable (at x within t)" |
|
295 by (metis DERIV_subset field_differentiable_def) |
|
296 |
|
297 lemma field_differentiable_at_within: |
|
298 "\<lbrakk>f field_differentiable (at x)\<rbrakk> |
|
299 \<Longrightarrow> f field_differentiable (at x within s)" |
|
300 unfolding field_differentiable_def |
|
301 by (metis DERIV_subset top_greatest) |
|
302 |
|
303 lemma field_differentiable_linear [simp,derivative_intros]: "(op * c) field_differentiable F" |
|
304 proof - |
|
305 show ?thesis |
|
306 unfolding field_differentiable_def has_field_derivative_def mult_commute_abs |
|
307 by (force intro: has_derivative_mult_right) |
|
308 qed |
|
309 |
|
310 lemma field_differentiable_const [simp,derivative_intros]: "(\<lambda>z. c) field_differentiable F" |
|
311 unfolding field_differentiable_def has_field_derivative_def |
|
312 by (rule exI [where x=0]) |
|
313 (metis has_derivative_const lambda_zero) |
|
314 |
|
315 lemma field_differentiable_ident [simp,derivative_intros]: "(\<lambda>z. z) field_differentiable F" |
|
316 unfolding field_differentiable_def has_field_derivative_def |
|
317 by (rule exI [where x=1]) |
|
318 (simp add: lambda_one [symmetric]) |
|
319 |
|
320 lemma field_differentiable_id [simp,derivative_intros]: "id field_differentiable F" |
|
321 unfolding id_def by (rule field_differentiable_ident) |
|
322 |
|
323 lemma field_differentiable_minus [derivative_intros]: |
|
324 "f field_differentiable F \<Longrightarrow> (\<lambda>z. - (f z)) field_differentiable F" |
|
325 unfolding field_differentiable_def |
|
326 by (metis field_differentiable_minus) |
|
327 |
|
328 lemma field_differentiable_add [derivative_intros]: |
|
329 assumes "f field_differentiable F" "g field_differentiable F" |
|
330 shows "(\<lambda>z. f z + g z) field_differentiable F" |
|
331 using assms unfolding field_differentiable_def |
|
332 by (metis field_differentiable_add) |
|
333 |
|
334 lemma field_differentiable_add_const [simp,derivative_intros]: |
|
335 "op + c field_differentiable F" |
|
336 by (simp add: field_differentiable_add) |
|
337 |
|
338 lemma field_differentiable_setsum [derivative_intros]: |
|
339 "(\<And>i. i \<in> I \<Longrightarrow> (f i) field_differentiable F) \<Longrightarrow> (\<lambda>z. \<Sum>i\<in>I. f i z) field_differentiable F" |
|
340 by (induct I rule: infinite_finite_induct) |
|
341 (auto intro: field_differentiable_add field_differentiable_const) |
|
342 |
|
343 lemma field_differentiable_diff [derivative_intros]: |
|
344 assumes "f field_differentiable F" "g field_differentiable F" |
|
345 shows "(\<lambda>z. f z - g z) field_differentiable F" |
|
346 using assms unfolding field_differentiable_def |
|
347 by (metis field_differentiable_diff) |
|
348 |
|
349 lemma field_differentiable_inverse [derivative_intros]: |
|
350 assumes "f field_differentiable (at a within s)" "f a \<noteq> 0" |
|
351 shows "(\<lambda>z. inverse (f z)) field_differentiable (at a within s)" |
|
352 using assms unfolding field_differentiable_def |
|
353 by (metis DERIV_inverse_fun) |
|
354 |
|
355 lemma field_differentiable_mult [derivative_intros]: |
|
356 assumes "f field_differentiable (at a within s)" |
|
357 "g field_differentiable (at a within s)" |
|
358 shows "(\<lambda>z. f z * g z) field_differentiable (at a within s)" |
|
359 using assms unfolding field_differentiable_def |
|
360 by (metis DERIV_mult [of f _ a s g]) |
|
361 |
|
362 lemma field_differentiable_divide [derivative_intros]: |
|
363 assumes "f field_differentiable (at a within s)" |
|
364 "g field_differentiable (at a within s)" |
|
365 "g a \<noteq> 0" |
|
366 shows "(\<lambda>z. f z / g z) field_differentiable (at a within s)" |
|
367 using assms unfolding field_differentiable_def |
|
368 by (metis DERIV_divide [of f _ a s g]) |
|
369 |
|
370 lemma field_differentiable_power [derivative_intros]: |
|
371 assumes "f field_differentiable (at a within s)" |
|
372 shows "(\<lambda>z. f z ^ n) field_differentiable (at a within s)" |
|
373 using assms unfolding field_differentiable_def |
|
374 by (metis DERIV_power) |
|
375 |
|
376 lemma field_differentiable_transform_within: |
|
377 "0 < d \<Longrightarrow> |
|
378 x \<in> s \<Longrightarrow> |
|
379 (\<And>x'. x' \<in> s \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x') \<Longrightarrow> |
|
380 f field_differentiable (at x within s) |
|
381 \<Longrightarrow> g field_differentiable (at x within s)" |
|
382 unfolding field_differentiable_def has_field_derivative_def |
|
383 by (blast intro: has_derivative_transform_within) |
|
384 |
|
385 lemma field_differentiable_compose_within: |
|
386 assumes "f field_differentiable (at a within s)" |
|
387 "g field_differentiable (at (f a) within f`s)" |
|
388 shows "(g o f) field_differentiable (at a within s)" |
|
389 using assms unfolding field_differentiable_def |
|
390 by (metis DERIV_image_chain) |
|
391 |
|
392 lemma field_differentiable_compose: |
|
393 "f field_differentiable at z \<Longrightarrow> g field_differentiable at (f z) |
|
394 \<Longrightarrow> (g o f) field_differentiable at z" |
|
395 by (metis field_differentiable_at_within field_differentiable_compose_within) |
|
396 |
|
397 lemma field_differentiable_within_open: |
|
398 "\<lbrakk>a \<in> s; open s\<rbrakk> \<Longrightarrow> f field_differentiable at a within s \<longleftrightarrow> |
|
399 f field_differentiable at a" |
|
400 unfolding field_differentiable_def |
|
401 by (metis at_within_open) |
|
402 |
|
403 subsection\<open>Caratheodory characterization\<close> |
|
404 |
|
405 lemma field_differentiable_caratheodory_at: |
|
406 "f field_differentiable (at z) \<longleftrightarrow> |
|
407 (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z) g)" |
|
408 using CARAT_DERIV [of f] |
|
409 by (simp add: field_differentiable_def has_field_derivative_def) |
|
410 |
|
411 lemma field_differentiable_caratheodory_within: |
|
412 "f field_differentiable (at z within s) \<longleftrightarrow> |
|
413 (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z within s) g)" |
|
414 using DERIV_caratheodory_within [of f] |
|
415 by (simp add: field_differentiable_def has_field_derivative_def) |
|
416 |
|
417 subsection\<open>Holomorphic\<close> |
|
418 |
|
419 definition holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool" |
|
420 (infixl "(holomorphic'_on)" 50) |
|
421 where "f holomorphic_on s \<equiv> \<forall>x\<in>s. f field_differentiable (at x within s)" |
|
422 |
|
423 named_theorems holomorphic_intros "structural introduction rules for holomorphic_on" |
|
424 |
|
425 lemma holomorphic_onI [intro?]: "(\<And>x. x \<in> s \<Longrightarrow> f field_differentiable (at x within s)) \<Longrightarrow> f holomorphic_on s" |
|
426 by (simp add: holomorphic_on_def) |
|
427 |
|
428 lemma holomorphic_onD [dest?]: "\<lbrakk>f holomorphic_on s; x \<in> s\<rbrakk> \<Longrightarrow> f field_differentiable (at x within s)" |
|
429 by (simp add: holomorphic_on_def) |
|
430 |
|
431 lemma holomorphic_on_imp_differentiable_at: |
|
432 "\<lbrakk>f holomorphic_on s; open s; x \<in> s\<rbrakk> \<Longrightarrow> f field_differentiable (at x)" |
|
433 using at_within_open holomorphic_on_def by fastforce |
|
434 |
|
435 lemma holomorphic_on_empty [holomorphic_intros]: "f holomorphic_on {}" |
|
436 by (simp add: holomorphic_on_def) |
|
437 |
|
438 lemma holomorphic_on_open: |
|
439 "open s \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>f'. DERIV f x :> f')" |
|
440 by (auto simp: holomorphic_on_def field_differentiable_def has_field_derivative_def at_within_open [of _ s]) |
|
441 |
|
442 lemma holomorphic_on_imp_continuous_on: |
|
443 "f holomorphic_on s \<Longrightarrow> continuous_on s f" |
|
444 by (metis field_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def) |
|
445 |
|
446 lemma holomorphic_on_subset [elim]: |
|
447 "f holomorphic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f holomorphic_on t" |
|
448 unfolding holomorphic_on_def |
|
449 by (metis field_differentiable_within_subset subsetD) |
|
450 |
|
451 lemma holomorphic_transform: "\<lbrakk>f holomorphic_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g holomorphic_on s" |
|
452 by (metis field_differentiable_transform_within linordered_field_no_ub holomorphic_on_def) |
|
453 |
|
454 lemma holomorphic_cong: "s = t ==> (\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> g holomorphic_on t" |
|
455 by (metis holomorphic_transform) |
|
456 |
|
457 lemma holomorphic_on_linear [simp, holomorphic_intros]: "(op * c) holomorphic_on s" |
|
458 unfolding holomorphic_on_def by (metis field_differentiable_linear) |
|
459 |
|
460 lemma holomorphic_on_const [simp, holomorphic_intros]: "(\<lambda>z. c) holomorphic_on s" |
|
461 unfolding holomorphic_on_def by (metis field_differentiable_const) |
|
462 |
|
463 lemma holomorphic_on_ident [simp, holomorphic_intros]: "(\<lambda>x. x) holomorphic_on s" |
|
464 unfolding holomorphic_on_def by (metis field_differentiable_ident) |
|
465 |
|
466 lemma holomorphic_on_id [simp, holomorphic_intros]: "id holomorphic_on s" |
|
467 unfolding id_def by (rule holomorphic_on_ident) |
|
468 |
|
469 lemma holomorphic_on_compose: |
|
470 "f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s) \<Longrightarrow> (g o f) holomorphic_on s" |
|
471 using field_differentiable_compose_within[of f _ s g] |
|
472 by (auto simp: holomorphic_on_def) |
|
473 |
|
474 lemma holomorphic_on_compose_gen: |
|
475 "f holomorphic_on s \<Longrightarrow> g holomorphic_on t \<Longrightarrow> f ` s \<subseteq> t \<Longrightarrow> (g o f) holomorphic_on s" |
|
476 by (metis holomorphic_on_compose holomorphic_on_subset) |
|
477 |
|
478 lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on s" |
|
479 by (metis field_differentiable_minus holomorphic_on_def) |
|
480 |
|
481 lemma holomorphic_on_add [holomorphic_intros]: |
|
482 "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) holomorphic_on s" |
|
483 unfolding holomorphic_on_def by (metis field_differentiable_add) |
|
484 |
|
485 lemma holomorphic_on_diff [holomorphic_intros]: |
|
486 "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) holomorphic_on s" |
|
487 unfolding holomorphic_on_def by (metis field_differentiable_diff) |
|
488 |
|
489 lemma holomorphic_on_mult [holomorphic_intros]: |
|
490 "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) holomorphic_on s" |
|
491 unfolding holomorphic_on_def by (metis field_differentiable_mult) |
|
492 |
|
493 lemma holomorphic_on_inverse [holomorphic_intros]: |
|
494 "\<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" |
|
495 unfolding holomorphic_on_def by (metis field_differentiable_inverse) |
|
496 |
|
497 lemma holomorphic_on_divide [holomorphic_intros]: |
|
498 "\<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" |
|
499 unfolding holomorphic_on_def by (metis field_differentiable_divide) |
|
500 |
|
501 lemma holomorphic_on_power [holomorphic_intros]: |
|
502 "f holomorphic_on s \<Longrightarrow> (\<lambda>z. (f z)^n) holomorphic_on s" |
|
503 unfolding holomorphic_on_def by (metis field_differentiable_power) |
|
504 |
|
505 lemma holomorphic_on_setsum [holomorphic_intros]: |
|
506 "(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) holomorphic_on s" |
|
507 unfolding holomorphic_on_def by (metis field_differentiable_setsum) |
|
508 |
|
509 lemma DERIV_deriv_iff_field_differentiable: |
|
510 "DERIV f x :> deriv f x \<longleftrightarrow> f field_differentiable at x" |
|
511 unfolding field_differentiable_def by (metis DERIV_imp_deriv) |
|
512 |
|
513 lemma holomorphic_derivI: |
|
514 "\<lbrakk>f holomorphic_on S; open S; x \<in> S\<rbrakk> |
|
515 \<Longrightarrow> (f has_field_derivative deriv f x) (at x within T)" |
|
516 by (metis DERIV_deriv_iff_field_differentiable at_within_open holomorphic_on_def has_field_derivative_at_within) |
|
517 |
|
518 lemma complex_derivative_chain: |
|
519 "f field_differentiable at x \<Longrightarrow> g field_differentiable at (f x) |
|
520 \<Longrightarrow> deriv (g o f) x = deriv g (f x) * deriv f x" |
|
521 by (metis DERIV_deriv_iff_field_differentiable DERIV_chain DERIV_imp_deriv) |
|
522 |
|
523 lemma deriv_linear [simp]: "deriv (\<lambda>w. c * w) = (\<lambda>z. c)" |
|
524 by (metis DERIV_imp_deriv DERIV_cmult_Id) |
|
525 |
|
526 lemma deriv_ident [simp]: "deriv (\<lambda>w. w) = (\<lambda>z. 1)" |
|
527 by (metis DERIV_imp_deriv DERIV_ident) |
|
528 |
|
529 lemma deriv_id [simp]: "deriv id = (\<lambda>z. 1)" |
|
530 by (simp add: id_def) |
|
531 |
|
532 lemma deriv_const [simp]: "deriv (\<lambda>w. c) = (\<lambda>z. 0)" |
|
533 by (metis DERIV_imp_deriv DERIV_const) |
|
534 |
|
535 lemma deriv_add [simp]: |
|
536 "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> |
|
537 \<Longrightarrow> deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z" |
|
538 unfolding DERIV_deriv_iff_field_differentiable[symmetric] |
|
539 by (auto intro!: DERIV_imp_deriv derivative_intros) |
|
540 |
|
541 lemma deriv_diff [simp]: |
|
542 "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> |
|
543 \<Longrightarrow> deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z" |
|
544 unfolding DERIV_deriv_iff_field_differentiable[symmetric] |
|
545 by (auto intro!: DERIV_imp_deriv derivative_intros) |
|
546 |
|
547 lemma deriv_mult [simp]: |
|
548 "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> |
|
549 \<Longrightarrow> deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z" |
|
550 unfolding DERIV_deriv_iff_field_differentiable[symmetric] |
|
551 by (auto intro!: DERIV_imp_deriv derivative_eq_intros) |
|
552 |
|
553 lemma deriv_cmult [simp]: |
|
554 "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. c * f w) z = c * deriv f z" |
|
555 unfolding DERIV_deriv_iff_field_differentiable[symmetric] |
|
556 by (auto intro!: DERIV_imp_deriv derivative_eq_intros) |
|
557 |
|
558 lemma deriv_cmult_right [simp]: |
|
559 "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w * c) z = deriv f z * c" |
|
560 unfolding DERIV_deriv_iff_field_differentiable[symmetric] |
|
561 by (auto intro!: DERIV_imp_deriv derivative_eq_intros) |
|
562 |
|
563 lemma deriv_cdivide_right [simp]: |
|
564 "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w / c) z = deriv f z / c" |
|
565 unfolding Fields.field_class.field_divide_inverse |
|
566 by (blast intro: deriv_cmult_right) |
|
567 |
|
568 lemma complex_derivative_transform_within_open: |
|
569 "\<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> |
|
570 \<Longrightarrow> deriv f z = deriv g z" |
|
571 unfolding holomorphic_on_def |
|
572 by (rule DERIV_imp_deriv) |
|
573 (metis DERIV_deriv_iff_field_differentiable DERIV_transform_within_open at_within_open) |
|
574 |
|
575 lemma deriv_compose_linear: |
|
576 "f field_differentiable at (c * z) \<Longrightarrow> deriv (\<lambda>w. f (c * w)) z = c * deriv f (c * z)" |
|
577 apply (rule DERIV_imp_deriv) |
|
578 apply (simp add: DERIV_deriv_iff_field_differentiable [symmetric]) |
|
579 apply (drule DERIV_chain' [of "times c" c z UNIV f "deriv f (c * z)", OF DERIV_cmult_Id]) |
|
580 apply (simp add: algebra_simps) |
|
581 done |
|
582 |
|
583 lemma nonzero_deriv_nonconstant: |
|
584 assumes df: "DERIV f \<xi> :> df" and S: "open S" "\<xi> \<in> S" and "df \<noteq> 0" |
|
585 shows "\<not> f constant_on S" |
|
586 unfolding constant_on_def |
|
587 by (metis \<open>df \<noteq> 0\<close> DERIV_transform_within_open [OF df S] DERIV_const DERIV_unique) |
|
588 |
|
589 lemma holomorphic_nonconstant: |
|
590 assumes holf: "f holomorphic_on S" and "open S" "\<xi> \<in> S" "deriv f \<xi> \<noteq> 0" |
|
591 shows "\<not> f constant_on S" |
|
592 apply (rule nonzero_deriv_nonconstant [of f "deriv f \<xi>" \<xi> S]) |
|
593 using assms |
|
594 apply (auto simp: holomorphic_derivI) |
|
595 done |
|
596 |
|
597 subsection\<open>Analyticity on a set\<close> |
|
598 |
|
599 definition analytic_on (infixl "(analytic'_on)" 50) |
|
600 where |
|
601 "f analytic_on s \<equiv> \<forall>x \<in> s. \<exists>e. 0 < e \<and> f holomorphic_on (ball x e)" |
|
602 |
|
603 lemma analytic_imp_holomorphic: "f analytic_on s \<Longrightarrow> f holomorphic_on s" |
|
604 by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def) |
|
605 (metis centre_in_ball field_differentiable_at_within) |
|
606 |
|
607 lemma analytic_on_open: "open s \<Longrightarrow> f analytic_on s \<longleftrightarrow> f holomorphic_on s" |
|
608 apply (auto simp: analytic_imp_holomorphic) |
|
609 apply (auto simp: analytic_on_def holomorphic_on_def) |
|
610 by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball) |
|
611 |
|
612 lemma analytic_on_imp_differentiable_at: |
|
613 "f analytic_on s \<Longrightarrow> x \<in> s \<Longrightarrow> f field_differentiable (at x)" |
|
614 apply (auto simp: analytic_on_def holomorphic_on_def) |
|
615 by (metis Topology_Euclidean_Space.open_ball centre_in_ball field_differentiable_within_open) |
|
616 |
|
617 lemma analytic_on_subset: "f analytic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f analytic_on t" |
|
618 by (auto simp: analytic_on_def) |
|
619 |
|
620 lemma analytic_on_Un: "f analytic_on (s \<union> t) \<longleftrightarrow> f analytic_on s \<and> f analytic_on t" |
|
621 by (auto simp: analytic_on_def) |
|
622 |
|
623 lemma analytic_on_Union: "f analytic_on (\<Union>s) \<longleftrightarrow> (\<forall>t \<in> s. f analytic_on t)" |
|
624 by (auto simp: analytic_on_def) |
|
625 |
|
626 lemma analytic_on_UN: "f analytic_on (\<Union>i\<in>I. s i) \<longleftrightarrow> (\<forall>i\<in>I. f analytic_on (s i))" |
|
627 by (auto simp: analytic_on_def) |
|
628 |
|
629 lemma analytic_on_holomorphic: |
|
630 "f analytic_on s \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f holomorphic_on t)" |
|
631 (is "?lhs = ?rhs") |
|
632 proof - |
|
633 have "?lhs \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t)" |
|
634 proof safe |
|
635 assume "f analytic_on s" |
|
636 then show "\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t" |
|
637 apply (simp add: analytic_on_def) |
|
638 apply (rule exI [where x="\<Union>{u. open u \<and> f analytic_on u}"], auto) |
|
639 apply (metis Topology_Euclidean_Space.open_ball analytic_on_open centre_in_ball) |
|
640 by (metis analytic_on_def) |
|
641 next |
|
642 fix t |
|
643 assume "open t" "s \<subseteq> t" "f analytic_on t" |
|
644 then show "f analytic_on s" |
|
645 by (metis analytic_on_subset) |
|
646 qed |
|
647 also have "... \<longleftrightarrow> ?rhs" |
|
648 by (auto simp: analytic_on_open) |
|
649 finally show ?thesis . |
|
650 qed |
|
651 |
|
652 lemma analytic_on_linear: "(op * c) analytic_on s" |
|
653 by (auto simp add: analytic_on_holomorphic holomorphic_on_linear) |
|
654 |
|
655 lemma analytic_on_const: "(\<lambda>z. c) analytic_on s" |
|
656 by (metis analytic_on_def holomorphic_on_const zero_less_one) |
|
657 |
|
658 lemma analytic_on_ident: "(\<lambda>x. x) analytic_on s" |
|
659 by (simp add: analytic_on_def holomorphic_on_ident gt_ex) |
|
660 |
|
661 lemma analytic_on_id: "id analytic_on s" |
|
662 unfolding id_def by (rule analytic_on_ident) |
|
663 |
|
664 lemma analytic_on_compose: |
|
665 assumes f: "f analytic_on s" |
|
666 and g: "g analytic_on (f ` s)" |
|
667 shows "(g o f) analytic_on s" |
|
668 unfolding analytic_on_def |
|
669 proof (intro ballI) |
|
670 fix x |
|
671 assume x: "x \<in> s" |
|
672 then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f |
|
673 by (metis analytic_on_def) |
|
674 obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g |
|
675 by (metis analytic_on_def g image_eqI x) |
|
676 have "isCont f x" |
|
677 by (metis analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at f x) |
|
678 with e' obtain d where d: "0 < d" and fd: "f ` ball x d \<subseteq> ball (f x) e'" |
|
679 by (auto simp: continuous_at_ball) |
|
680 have "g \<circ> f holomorphic_on ball x (min d e)" |
|
681 apply (rule holomorphic_on_compose) |
|
682 apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) |
|
683 by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball) |
|
684 then show "\<exists>e>0. g \<circ> f holomorphic_on ball x e" |
|
685 by (metis d e min_less_iff_conj) |
|
686 qed |
|
687 |
|
688 lemma analytic_on_compose_gen: |
|
689 "f analytic_on s \<Longrightarrow> g analytic_on t \<Longrightarrow> (\<And>z. z \<in> s \<Longrightarrow> f z \<in> t) |
|
690 \<Longrightarrow> g o f analytic_on s" |
|
691 by (metis analytic_on_compose analytic_on_subset image_subset_iff) |
|
692 |
|
693 lemma analytic_on_neg: |
|
694 "f analytic_on s \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on s" |
|
695 by (metis analytic_on_holomorphic holomorphic_on_minus) |
|
696 |
|
697 lemma analytic_on_add: |
|
698 assumes f: "f analytic_on s" |
|
699 and g: "g analytic_on s" |
|
700 shows "(\<lambda>z. f z + g z) analytic_on s" |
|
701 unfolding analytic_on_def |
|
702 proof (intro ballI) |
|
703 fix z |
|
704 assume z: "z \<in> s" |
|
705 then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f |
|
706 by (metis analytic_on_def) |
|
707 obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g |
|
708 by (metis analytic_on_def g z) |
|
709 have "(\<lambda>z. f z + g z) holomorphic_on ball z (min e e')" |
|
710 apply (rule holomorphic_on_add) |
|
711 apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) |
|
712 by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) |
|
713 then show "\<exists>e>0. (\<lambda>z. f z + g z) holomorphic_on ball z e" |
|
714 by (metis e e' min_less_iff_conj) |
|
715 qed |
|
716 |
|
717 lemma analytic_on_diff: |
|
718 assumes f: "f analytic_on s" |
|
719 and g: "g analytic_on s" |
|
720 shows "(\<lambda>z. f z - g z) analytic_on s" |
|
721 unfolding analytic_on_def |
|
722 proof (intro ballI) |
|
723 fix z |
|
724 assume z: "z \<in> s" |
|
725 then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f |
|
726 by (metis analytic_on_def) |
|
727 obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g |
|
728 by (metis analytic_on_def g z) |
|
729 have "(\<lambda>z. f z - g z) holomorphic_on ball z (min e e')" |
|
730 apply (rule holomorphic_on_diff) |
|
731 apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) |
|
732 by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) |
|
733 then show "\<exists>e>0. (\<lambda>z. f z - g z) holomorphic_on ball z e" |
|
734 by (metis e e' min_less_iff_conj) |
|
735 qed |
|
736 |
|
737 lemma analytic_on_mult: |
|
738 assumes f: "f analytic_on s" |
|
739 and g: "g analytic_on s" |
|
740 shows "(\<lambda>z. f z * g z) analytic_on s" |
|
741 unfolding analytic_on_def |
|
742 proof (intro ballI) |
|
743 fix z |
|
744 assume z: "z \<in> s" |
|
745 then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f |
|
746 by (metis analytic_on_def) |
|
747 obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g |
|
748 by (metis analytic_on_def g z) |
|
749 have "(\<lambda>z. f z * g z) holomorphic_on ball z (min e e')" |
|
750 apply (rule holomorphic_on_mult) |
|
751 apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) |
|
752 by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) |
|
753 then show "\<exists>e>0. (\<lambda>z. f z * g z) holomorphic_on ball z e" |
|
754 by (metis e e' min_less_iff_conj) |
|
755 qed |
|
756 |
|
757 lemma analytic_on_inverse: |
|
758 assumes f: "f analytic_on s" |
|
759 and nz: "(\<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0)" |
|
760 shows "(\<lambda>z. inverse (f z)) analytic_on s" |
|
761 unfolding analytic_on_def |
|
762 proof (intro ballI) |
|
763 fix z |
|
764 assume z: "z \<in> s" |
|
765 then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f |
|
766 by (metis analytic_on_def) |
|
767 have "continuous_on (ball z e) f" |
|
768 by (metis fh holomorphic_on_imp_continuous_on) |
|
769 then obtain e' where e': "0 < e'" and nz': "\<And>y. dist z y < e' \<Longrightarrow> f y \<noteq> 0" |
|
770 by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz) |
|
771 have "(\<lambda>z. inverse (f z)) holomorphic_on ball z (min e e')" |
|
772 apply (rule holomorphic_on_inverse) |
|
773 apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball) |
|
774 by (metis nz' mem_ball min_less_iff_conj) |
|
775 then show "\<exists>e>0. (\<lambda>z. inverse (f z)) holomorphic_on ball z e" |
|
776 by (metis e e' min_less_iff_conj) |
|
777 qed |
|
778 |
|
779 lemma analytic_on_divide: |
|
780 assumes f: "f analytic_on s" |
|
781 and g: "g analytic_on s" |
|
782 and nz: "(\<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0)" |
|
783 shows "(\<lambda>z. f z / g z) analytic_on s" |
|
784 unfolding divide_inverse |
|
785 by (metis analytic_on_inverse analytic_on_mult f g nz) |
|
786 |
|
787 lemma analytic_on_power: |
|
788 "f analytic_on s \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on s" |
|
789 by (induct n) (auto simp: analytic_on_const analytic_on_mult) |
|
790 |
|
791 lemma analytic_on_setsum: |
|
792 "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) analytic_on s" |
|
793 by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add) |
|
794 |
|
795 lemma deriv_left_inverse: |
|
796 assumes "f holomorphic_on S" and "g holomorphic_on T" |
|
797 and "open S" and "open T" |
|
798 and "f ` S \<subseteq> T" |
|
799 and [simp]: "\<And>z. z \<in> S \<Longrightarrow> g (f z) = z" |
|
800 and "w \<in> S" |
|
801 shows "deriv f w * deriv g (f w) = 1" |
|
802 proof - |
|
803 have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w" |
|
804 by (simp add: algebra_simps) |
|
805 also have "... = deriv (g o f) w" |
|
806 using assms |
|
807 by (metis analytic_on_imp_differentiable_at analytic_on_open complex_derivative_chain image_subset_iff) |
|
808 also have "... = deriv id w" |
|
809 apply (rule complex_derivative_transform_within_open [where s=S]) |
|
810 apply (rule assms holomorphic_on_compose_gen holomorphic_intros)+ |
|
811 apply simp |
|
812 done |
|
813 also have "... = 1" |
|
814 by simp |
|
815 finally show ?thesis . |
|
816 qed |
|
817 |
|
818 subsection\<open>analyticity at a point\<close> |
|
819 |
|
820 lemma analytic_at_ball: |
|
821 "f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)" |
|
822 by (metis analytic_on_def singleton_iff) |
|
823 |
|
824 lemma analytic_at: |
|
825 "f analytic_on {z} \<longleftrightarrow> (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s)" |
|
826 by (metis analytic_on_holomorphic empty_subsetI insert_subset) |
|
827 |
|
828 lemma analytic_on_analytic_at: |
|
829 "f analytic_on s \<longleftrightarrow> (\<forall>z \<in> s. f analytic_on {z})" |
|
830 by (metis analytic_at_ball analytic_on_def) |
|
831 |
|
832 lemma analytic_at_two: |
|
833 "f analytic_on {z} \<and> g analytic_on {z} \<longleftrightarrow> |
|
834 (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s \<and> g holomorphic_on s)" |
|
835 (is "?lhs = ?rhs") |
|
836 proof |
|
837 assume ?lhs |
|
838 then obtain s t |
|
839 where st: "open s" "z \<in> s" "f holomorphic_on s" |
|
840 "open t" "z \<in> t" "g holomorphic_on t" |
|
841 by (auto simp: analytic_at) |
|
842 show ?rhs |
|
843 apply (rule_tac x="s \<inter> t" in exI) |
|
844 using st |
|
845 apply (auto simp: Diff_subset holomorphic_on_subset) |
|
846 done |
|
847 next |
|
848 assume ?rhs |
|
849 then show ?lhs |
|
850 by (force simp add: analytic_at) |
|
851 qed |
|
852 |
|
853 subsection\<open>Combining theorems for derivative with ``analytic at'' hypotheses\<close> |
|
854 |
|
855 lemma |
|
856 assumes "f analytic_on {z}" "g analytic_on {z}" |
|
857 shows complex_derivative_add_at: "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z" |
|
858 and complex_derivative_diff_at: "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z" |
|
859 and complex_derivative_mult_at: "deriv (\<lambda>w. f w * g w) z = |
|
860 f z * deriv g z + deriv f z * g z" |
|
861 proof - |
|
862 obtain s where s: "open s" "z \<in> s" "f holomorphic_on s" "g holomorphic_on s" |
|
863 using assms by (metis analytic_at_two) |
|
864 show "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z" |
|
865 apply (rule DERIV_imp_deriv [OF DERIV_add]) |
|
866 using s |
|
867 apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) |
|
868 done |
|
869 show "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z" |
|
870 apply (rule DERIV_imp_deriv [OF DERIV_diff]) |
|
871 using s |
|
872 apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) |
|
873 done |
|
874 show "deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z" |
|
875 apply (rule DERIV_imp_deriv [OF DERIV_mult']) |
|
876 using s |
|
877 apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) |
|
878 done |
|
879 qed |
|
880 |
|
881 lemma deriv_cmult_at: |
|
882 "f analytic_on {z} \<Longrightarrow> deriv (\<lambda>w. c * f w) z = c * deriv f z" |
|
883 by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) |
|
884 |
|
885 lemma deriv_cmult_right_at: |
|
886 "f analytic_on {z} \<Longrightarrow> deriv (\<lambda>w. f w * c) z = deriv f z * c" |
|
887 by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) |
|
888 |
|
889 subsection\<open>Complex differentiation of sequences and series\<close> |
|
890 |
|
891 (* TODO: Could probably be simplified using Uniform_Limit *) |
|
892 lemma has_complex_derivative_sequence: |
|
893 fixes s :: "complex set" |
|
894 assumes cvs: "convex s" |
|
895 and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" |
|
896 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" |
|
897 and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially" |
|
898 shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> |
|
899 (g has_field_derivative (g' x)) (at x within s)" |
|
900 proof - |
|
901 from assms obtain x l where x: "x \<in> s" and tf: "((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially" |
|
902 by blast |
|
903 { fix e::real assume e: "e > 0" |
|
904 then obtain N where N: "\<forall>n\<ge>N. \<forall>x. x \<in> s \<longrightarrow> cmod (f' n x - g' x) \<le> e" |
|
905 by (metis conv) |
|
906 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" |
|
907 proof (rule exI [of _ N], clarify) |
|
908 fix n y h |
|
909 assume "N \<le> n" "y \<in> s" |
|
910 then have "cmod (f' n y - g' y) \<le> e" |
|
911 by (metis N) |
|
912 then have "cmod h * cmod (f' n y - g' y) \<le> cmod h * e" |
|
913 by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) |
|
914 then show "cmod (f' n y * h - g' y * h) \<le> e * cmod h" |
|
915 by (simp add: norm_mult [symmetric] field_simps) |
|
916 qed |
|
917 } note ** = this |
|
918 show ?thesis |
|
919 unfolding has_field_derivative_def |
|
920 proof (rule has_derivative_sequence [OF cvs _ _ x]) |
|
921 show "\<forall>n. \<forall>x\<in>s. (f n has_derivative (op * (f' n x))) (at x within s)" |
|
922 by (metis has_field_derivative_def df) |
|
923 next show "(\<lambda>n. f n x) \<longlonglongrightarrow> l" |
|
924 by (rule tf) |
|
925 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" |
|
926 by (blast intro: **) |
|
927 qed |
|
928 qed |
|
929 |
|
930 lemma has_complex_derivative_series: |
|
931 fixes s :: "complex set" |
|
932 assumes cvs: "convex s" |
|
933 and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" |
|
934 and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s |
|
935 \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e" |
|
936 and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) sums l)" |
|
937 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))" |
|
938 proof - |
|
939 from assms obtain x l where x: "x \<in> s" and sf: "((\<lambda>n. f n x) sums l)" |
|
940 by blast |
|
941 { fix e::real assume e: "e > 0" |
|
942 then obtain N where N: "\<forall>n x. n \<ge> N \<longrightarrow> x \<in> s |
|
943 \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e" |
|
944 by (metis conv) |
|
945 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" |
|
946 proof (rule exI [of _ N], clarify) |
|
947 fix n y h |
|
948 assume "N \<le> n" "y \<in> s" |
|
949 then have "cmod ((\<Sum>i<n. f' i y) - g' y) \<le> e" |
|
950 by (metis N) |
|
951 then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e" |
|
952 by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) |
|
953 then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h" |
|
954 by (simp add: norm_mult [symmetric] field_simps setsum_right_distrib) |
|
955 qed |
|
956 } note ** = this |
|
957 show ?thesis |
|
958 unfolding has_field_derivative_def |
|
959 proof (rule has_derivative_series [OF cvs _ _ x]) |
|
960 fix n x |
|
961 assume "x \<in> s" |
|
962 then show "((f n) has_derivative (\<lambda>z. z * f' n x)) (at x within s)" |
|
963 by (metis df has_field_derivative_def mult_commute_abs) |
|
964 next show " ((\<lambda>n. f n x) sums l)" |
|
965 by (rule sf) |
|
966 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" |
|
967 by (blast intro: **) |
|
968 qed |
|
969 qed |
|
970 |
|
971 |
|
972 lemma field_differentiable_series: |
|
973 fixes f :: "nat \<Rightarrow> complex \<Rightarrow> complex" |
|
974 assumes "convex s" "open s" |
|
975 assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" |
|
976 assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" |
|
977 assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" and x: "x \<in> s" |
|
978 shows "summable (\<lambda>n. f n x)" and "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x)" |
|
979 proof - |
|
980 from assms(4) obtain g' where A: "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" |
|
981 unfolding uniformly_convergent_on_def by blast |
|
982 from x and \<open>open s\<close> have s: "at x within s = at x" by (rule at_within_open) |
|
983 have "\<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)" |
|
984 by (intro has_field_derivative_series[of s f f' g' x0] assms A has_field_derivative_at_within) |
|
985 then obtain g where g: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x" |
|
986 "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast |
|
987 from g[OF x] show "summable (\<lambda>n. f n x)" by (auto simp: summable_def) |
|
988 from g(2)[OF x] have g': "(g has_derivative op * (g' x)) (at x)" |
|
989 by (simp add: has_field_derivative_def s) |
|
990 have "((\<lambda>x. \<Sum>n. f n x) has_derivative op * (g' x)) (at x)" |
|
991 by (rule has_derivative_transform_within_open[OF g' \<open>open s\<close> x]) |
|
992 (insert g, auto simp: sums_iff) |
|
993 thus "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x)" unfolding differentiable_def |
|
994 by (auto simp: summable_def field_differentiable_def has_field_derivative_def) |
|
995 qed |
|
996 |
|
997 lemma field_differentiable_series': |
|
998 fixes f :: "nat \<Rightarrow> complex \<Rightarrow> complex" |
|
999 assumes "convex s" "open s" |
|
1000 assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" |
|
1001 assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" |
|
1002 assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" |
|
1003 shows "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x0)" |
|
1004 using field_differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+ |
|
1005 |
|
1006 subsection\<open>Bound theorem\<close> |
|
1007 |
|
1008 lemma field_differentiable_bound: |
|
1009 fixes s :: "complex set" |
|
1010 assumes cvs: "convex s" |
|
1011 and df: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z within s)" |
|
1012 and dn: "\<And>z. z \<in> s \<Longrightarrow> norm (f' z) \<le> B" |
|
1013 and "x \<in> s" "y \<in> s" |
|
1014 shows "norm(f x - f y) \<le> B * norm(x - y)" |
|
1015 apply (rule differentiable_bound [OF cvs]) |
|
1016 apply (rule ballI, erule df [unfolded has_field_derivative_def]) |
|
1017 apply (rule ballI, rule onorm_le, simp add: norm_mult mult_right_mono dn) |
|
1018 apply fact |
|
1019 apply fact |
|
1020 done |
|
1021 |
|
1022 subsection\<open>Inverse function theorem for complex derivatives\<close> |
|
1023 |
|
1024 lemma has_complex_derivative_inverse_basic: |
|
1025 fixes f :: "complex \<Rightarrow> complex" |
|
1026 shows "DERIV f (g y) :> f' \<Longrightarrow> |
|
1027 f' \<noteq> 0 \<Longrightarrow> |
|
1028 continuous (at y) g \<Longrightarrow> |
|
1029 open t \<Longrightarrow> |
|
1030 y \<in> t \<Longrightarrow> |
|
1031 (\<And>z. z \<in> t \<Longrightarrow> f (g z) = z) |
|
1032 \<Longrightarrow> DERIV g y :> inverse (f')" |
|
1033 unfolding has_field_derivative_def |
|
1034 apply (rule has_derivative_inverse_basic) |
|
1035 apply (auto simp: bounded_linear_mult_right) |
|
1036 done |
|
1037 |
|
1038 (*Used only once, in Multivariate/cauchy.ml. *) |
|
1039 lemma has_complex_derivative_inverse_strong: |
|
1040 fixes f :: "complex \<Rightarrow> complex" |
|
1041 shows "DERIV f x :> f' \<Longrightarrow> |
|
1042 f' \<noteq> 0 \<Longrightarrow> |
|
1043 open s \<Longrightarrow> |
|
1044 x \<in> s \<Longrightarrow> |
|
1045 continuous_on s f \<Longrightarrow> |
|
1046 (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z) |
|
1047 \<Longrightarrow> DERIV g (f x) :> inverse (f')" |
|
1048 unfolding has_field_derivative_def |
|
1049 apply (rule has_derivative_inverse_strong [of s x f g ]) |
|
1050 by auto |
|
1051 |
|
1052 lemma has_complex_derivative_inverse_strong_x: |
|
1053 fixes f :: "complex \<Rightarrow> complex" |
|
1054 shows "DERIV f (g y) :> f' \<Longrightarrow> |
|
1055 f' \<noteq> 0 \<Longrightarrow> |
|
1056 open s \<Longrightarrow> |
|
1057 continuous_on s f \<Longrightarrow> |
|
1058 g y \<in> s \<Longrightarrow> f(g y) = y \<Longrightarrow> |
|
1059 (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z) |
|
1060 \<Longrightarrow> DERIV g y :> inverse (f')" |
|
1061 unfolding has_field_derivative_def |
|
1062 apply (rule has_derivative_inverse_strong_x [of s g y f]) |
|
1063 by auto |
|
1064 |
|
1065 subsection \<open>Taylor on Complex Numbers\<close> |
|
1066 |
|
1067 lemma setsum_Suc_reindex: |
|
1068 fixes f :: "nat \<Rightarrow> 'a::ab_group_add" |
|
1069 shows "setsum f {0..n} = f 0 - f (Suc n) + setsum (\<lambda>i. f (Suc i)) {0..n}" |
|
1070 by (induct n) auto |
|
1071 |
|
1072 lemma complex_taylor: |
|
1073 assumes s: "convex s" |
|
1074 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)" |
|
1075 and B: "\<And>x. x \<in> s \<Longrightarrow> cmod (f (Suc n) x) \<le> B" |
|
1076 and w: "w \<in> s" |
|
1077 and z: "z \<in> s" |
|
1078 shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i))) |
|
1079 \<le> B * cmod(z - w)^(Suc n) / fact n" |
|
1080 proof - |
|
1081 have wzs: "closed_segment w z \<subseteq> s" using assms |
|
1082 by (metis convex_contains_segment) |
|
1083 { fix u |
|
1084 assume "u \<in> closed_segment w z" |
|
1085 then have "u \<in> s" |
|
1086 by (metis wzs subsetD) |
|
1087 have "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) + |
|
1088 f (Suc i) u * (z-u)^i / (fact i)) = |
|
1089 f (Suc n) u * (z-u) ^ n / (fact n)" |
|
1090 proof (induction n) |
|
1091 case 0 show ?case by simp |
|
1092 next |
|
1093 case (Suc n) |
|
1094 have "(\<Sum>i\<le>Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / (fact i) + |
|
1095 f (Suc i) u * (z-u) ^ i / (fact i)) = |
|
1096 f (Suc n) u * (z-u) ^ n / (fact n) + |
|
1097 f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) - |
|
1098 f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))" |
|
1099 using Suc by simp |
|
1100 also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))" |
|
1101 proof - |
|
1102 have "(fact(Suc n)) * |
|
1103 (f(Suc n) u *(z-u) ^ n / (fact n) + |
|
1104 f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / (fact(Suc n)) - |
|
1105 f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / (fact(Suc n))) = |
|
1106 ((fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / (fact n) + |
|
1107 ((fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / (fact(Suc n))) - |
|
1108 ((fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / (fact(Suc n))" |
|
1109 by (simp add: algebra_simps del: fact_Suc) |
|
1110 also have "... = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) + |
|
1111 (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - |
|
1112 (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" |
|
1113 by (simp del: fact_Suc) |
|
1114 also have "... = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) + |
|
1115 (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - |
|
1116 (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" |
|
1117 by (simp only: fact_Suc of_nat_mult ac_simps) simp |
|
1118 also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)" |
|
1119 by (simp add: algebra_simps) |
|
1120 finally show ?thesis |
|
1121 by (simp add: mult_left_cancel [where c = "(fact (Suc n))", THEN iffD1] del: fact_Suc) |
|
1122 qed |
|
1123 finally show ?case . |
|
1124 qed |
|
1125 then have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / (fact i))) |
|
1126 has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n)) |
|
1127 (at u within s)" |
|
1128 apply (intro derivative_eq_intros) |
|
1129 apply (blast intro: assms \<open>u \<in> s\<close>) |
|
1130 apply (rule refl)+ |
|
1131 apply (auto simp: field_simps) |
|
1132 done |
|
1133 } note sum_deriv = this |
|
1134 { fix u |
|
1135 assume u: "u \<in> closed_segment w z" |
|
1136 then have us: "u \<in> s" |
|
1137 by (metis wzs subsetD) |
|
1138 have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> cmod (f (Suc n) u) * cmod (u - z) ^ n" |
|
1139 by (metis norm_minus_commute order_refl) |
|
1140 also have "... \<le> cmod (f (Suc n) u) * cmod (z - w) ^ n" |
|
1141 by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u]) |
|
1142 also have "... \<le> B * cmod (z - w) ^ n" |
|
1143 by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us]) |
|
1144 finally have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> B * cmod (z - w) ^ n" . |
|
1145 } note cmod_bound = this |
|
1146 have "(\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)) = (\<Sum>i\<le>n. (f i z / (fact i)) * 0 ^ i)" |
|
1147 by simp |
|
1148 also have "\<dots> = f 0 z / (fact 0)" |
|
1149 by (subst setsum_zero_power) simp |
|
1150 finally have "cmod (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i))) |
|
1151 \<le> cmod ((\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)) - |
|
1152 (\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)))" |
|
1153 by (simp add: norm_minus_commute) |
|
1154 also have "... \<le> B * cmod (z - w) ^ n / (fact n) * cmod (w - z)" |
|
1155 apply (rule field_differentiable_bound |
|
1156 [where f' = "\<lambda>w. f (Suc n) w * (z - w)^n / (fact n)" |
|
1157 and s = "closed_segment w z", OF convex_closed_segment]) |
|
1158 apply (auto simp: ends_in_segment DERIV_subset [OF sum_deriv wzs] |
|
1159 norm_divide norm_mult norm_power divide_le_cancel cmod_bound) |
|
1160 done |
|
1161 also have "... \<le> B * cmod (z - w) ^ Suc n / (fact n)" |
|
1162 by (simp add: algebra_simps norm_minus_commute) |
|
1163 finally show ?thesis . |
|
1164 qed |
|
1165 |
|
1166 text\<open>Something more like the traditional MVT for real components\<close> |
|
1167 |
|
1168 lemma complex_mvt_line: |
|
1169 assumes "\<And>u. u \<in> closed_segment w z \<Longrightarrow> (f has_field_derivative f'(u)) (at u)" |
|
1170 shows "\<exists>u. u \<in> closed_segment w z \<and> Re(f z) - Re(f w) = Re(f'(u) * (z - w))" |
|
1171 proof - |
|
1172 have twz: "\<And>t. (1 - t) *\<^sub>R w + t *\<^sub>R z = w + t *\<^sub>R (z - w)" |
|
1173 by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib) |
|
1174 note assms[unfolded has_field_derivative_def, derivative_intros] |
|
1175 show ?thesis |
|
1176 apply (cut_tac mvt_simple |
|
1177 [of 0 1 "Re o f o (\<lambda>t. (1 - t) *\<^sub>R w + t *\<^sub>R z)" |
|
1178 "\<lambda>u. Re o (\<lambda>h. f'((1 - u) *\<^sub>R w + u *\<^sub>R z) * h) o (\<lambda>t. t *\<^sub>R (z - w))"]) |
|
1179 apply auto |
|
1180 apply (rule_tac x="(1 - x) *\<^sub>R w + x *\<^sub>R z" in exI) |
|
1181 apply (auto simp: closed_segment_def twz) [] |
|
1182 apply (intro derivative_eq_intros has_derivative_at_within, simp_all) |
|
1183 apply (simp add: fun_eq_iff real_vector.scale_right_diff_distrib) |
|
1184 apply (force simp: twz closed_segment_def) |
|
1185 done |
|
1186 qed |
|
1187 |
|
1188 lemma complex_taylor_mvt: |
|
1189 assumes "\<And>i x. \<lbrakk>x \<in> closed_segment w z; i \<le> n\<rbrakk> \<Longrightarrow> ((f i) has_field_derivative f (Suc i) x) (at x)" |
|
1190 shows "\<exists>u. u \<in> closed_segment w z \<and> |
|
1191 Re (f 0 z) = |
|
1192 Re ((\<Sum>i = 0..n. f i w * (z - w) ^ i / (fact i)) + |
|
1193 (f (Suc n) u * (z-u)^n / (fact n)) * (z - w))" |
|
1194 proof - |
|
1195 { fix u |
|
1196 assume u: "u \<in> closed_segment w z" |
|
1197 have "(\<Sum>i = 0..n. |
|
1198 (f (Suc i) u * (z-u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / |
|
1199 (fact i)) = |
|
1200 f (Suc 0) u - |
|
1201 (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / |
|
1202 (fact (Suc n)) + |
|
1203 (\<Sum>i = 0..n. |
|
1204 (f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) / |
|
1205 (fact (Suc i)))" |
|
1206 by (subst setsum_Suc_reindex) simp |
|
1207 also have "... = f (Suc 0) u - |
|
1208 (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / |
|
1209 (fact (Suc n)) + |
|
1210 (\<Sum>i = 0..n. |
|
1211 f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i)) - |
|
1212 f (Suc i) u * (z-u) ^ i / (fact i))" |
|
1213 by (simp only: diff_divide_distrib fact_cancel ac_simps) |
|
1214 also have "... = f (Suc 0) u - |
|
1215 (f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) / |
|
1216 (fact (Suc n)) + |
|
1217 f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u" |
|
1218 by (subst setsum_Suc_diff) auto |
|
1219 also have "... = f (Suc n) u * (z-u) ^ n / (fact n)" |
|
1220 by (simp only: algebra_simps diff_divide_distrib fact_cancel) |
|
1221 finally have "(\<Sum>i = 0..n. (f (Suc i) u * (z - u) ^ i |
|
1222 - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) = |
|
1223 f (Suc n) u * (z - u) ^ n / (fact n)" . |
|
1224 then have "((\<lambda>u. \<Sum>i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative |
|
1225 f (Suc n) u * (z - u) ^ n / (fact n)) (at u)" |
|
1226 apply (intro derivative_eq_intros)+ |
|
1227 apply (force intro: u assms) |
|
1228 apply (rule refl)+ |
|
1229 apply (auto simp: ac_simps) |
|
1230 done |
|
1231 } |
|
1232 then show ?thesis |
|
1233 apply (cut_tac complex_mvt_line [of w z "\<lambda>u. \<Sum>i = 0..n. f i u * (z-u) ^ i / (fact i)" |
|
1234 "\<lambda>u. (f (Suc n) u * (z-u)^n / (fact n))"]) |
|
1235 apply (auto simp add: intro: open_closed_segment) |
|
1236 done |
|
1237 qed |
|
1238 |
|
1239 |
|
1240 subsection \<open>Polynomal function extremal theorem, from HOL Light\<close> |
|
1241 |
|
1242 lemma polyfun_extremal_lemma: (*COMPLEX_POLYFUN_EXTREMAL_LEMMA in HOL Light*) |
|
1243 fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" |
|
1244 assumes "0 < e" |
|
1245 shows "\<exists>M. \<forall>z. M \<le> norm(z) \<longrightarrow> norm (\<Sum>i\<le>n. c(i) * z^i) \<le> e * norm(z) ^ (Suc n)" |
|
1246 proof (induct n) |
|
1247 case 0 with assms |
|
1248 show ?case |
|
1249 apply (rule_tac x="norm (c 0) / e" in exI) |
|
1250 apply (auto simp: field_simps) |
|
1251 done |
|
1252 next |
|
1253 case (Suc n) |
|
1254 obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" |
|
1255 using Suc assms by blast |
|
1256 show ?case |
|
1257 proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc) |
|
1258 fix z::'a |
|
1259 assume z1: "M \<le> norm z" and "1 + norm (c (Suc n)) / e \<le> norm z" |
|
1260 then have z2: "e + norm (c (Suc n)) \<le> e * norm z" |
|
1261 using assms by (simp add: field_simps) |
|
1262 have "norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" |
|
1263 using M [OF z1] by simp |
|
1264 then have "norm (\<Sum>i\<le>n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" |
|
1265 by simp |
|
1266 then have "norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" |
|
1267 by (blast intro: norm_triangle_le elim: ) |
|
1268 also have "... \<le> (e + norm (c (Suc n))) * norm z ^ Suc n" |
|
1269 by (simp add: norm_power norm_mult algebra_simps) |
|
1270 also have "... \<le> (e * norm z) * norm z ^ Suc n" |
|
1271 by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power) |
|
1272 finally show "norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc (Suc n)" |
|
1273 by simp |
|
1274 qed |
|
1275 qed |
|
1276 |
|
1277 lemma polyfun_extremal: (*COMPLEX_POLYFUN_EXTREMAL in HOL Light*) |
|
1278 fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" |
|
1279 assumes k: "c k \<noteq> 0" "1\<le>k" and kn: "k\<le>n" |
|
1280 shows "eventually (\<lambda>z. norm (\<Sum>i\<le>n. c(i) * z^i) \<ge> B) at_infinity" |
|
1281 using kn |
|
1282 proof (induction n) |
|
1283 case 0 |
|
1284 then show ?case |
|
1285 using k by simp |
|
1286 next |
|
1287 case (Suc m) |
|
1288 let ?even = ?case |
|
1289 show ?even |
|
1290 proof (cases "c (Suc m) = 0") |
|
1291 case True |
|
1292 then show ?even using Suc k |
|
1293 by auto (metis antisym_conv less_eq_Suc_le not_le) |
|
1294 next |
|
1295 case False |
|
1296 then obtain M where M: |
|
1297 "\<And>z. M \<le> norm z \<Longrightarrow> norm (\<Sum>i\<le>m. c i * z^i) \<le> norm (c (Suc m)) / 2 * norm z ^ Suc m" |
|
1298 using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc |
|
1299 by auto |
|
1300 have "\<exists>b. \<forall>z. b \<le> norm z \<longrightarrow> B \<le> norm (\<Sum>i\<le>Suc m. c i * z^i)" |
|
1301 proof (rule exI [where x="max M (max 1 (\<bar>B\<bar> / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc) |
|
1302 fix z::'a |
|
1303 assume z1: "M \<le> norm z" "1 \<le> norm z" |
|
1304 and "\<bar>B\<bar> * 2 / norm (c (Suc m)) \<le> norm z" |
|
1305 then have z2: "\<bar>B\<bar> \<le> norm (c (Suc m)) * norm z / 2" |
|
1306 using False by (simp add: field_simps) |
|
1307 have nz: "norm z \<le> norm z ^ Suc m" |
|
1308 by (metis \<open>1 \<le> norm z\<close> One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc) |
|
1309 have *: "\<And>y x. norm (c (Suc m)) * norm z / 2 \<le> norm y - norm x \<Longrightarrow> B \<le> norm (x + y)" |
|
1310 by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2) |
|
1311 have "norm z * norm (c (Suc m)) + 2 * norm (\<Sum>i\<le>m. c i * z^i) |
|
1312 \<le> norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m" |
|
1313 using M [of z] Suc z1 by auto |
|
1314 also have "... \<le> 2 * (norm (c (Suc m)) * norm z ^ Suc m)" |
|
1315 using nz by (simp add: mult_mono del: power_Suc) |
|
1316 finally show "B \<le> norm ((\<Sum>i\<le>m. c i * z^i) + c (Suc m) * z ^ Suc m)" |
|
1317 using Suc.IH |
|
1318 apply (auto simp: eventually_at_infinity) |
|
1319 apply (rule *) |
|
1320 apply (simp add: field_simps norm_mult norm_power) |
|
1321 done |
|
1322 qed |
|
1323 then show ?even |
|
1324 by (simp add: eventually_at_infinity) |
|
1325 qed |
|
1326 qed |
|
1327 |
|
1328 end |