author | paulson <lp15@cam.ac.uk> |
Wed, 22 Feb 2017 15:04:59 +0000 | |
changeset 65039 | 87972e6177bc |
parent 65036 | ab7e11730ad8 |
child 65064 | a4abec71279a |
permissions | -rw-r--r-- |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1 |
section \<open>Conformal Mappings. Consequences of Cauchy's integral theorem.\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
3 |
text\<open>By John Harrison et al. Ported from HOL Light by L C Paulson (2016)\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
4 |
|
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
5 |
text\<open>Also Cauchy's residue theorem by Wenda Li (2016)\<close> |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
6 |
|
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
7 |
theory Conformal_Mappings |
65039
87972e6177bc
New theory about Winding Numbers
paulson <lp15@cam.ac.uk>
parents:
65036
diff
changeset
|
8 |
imports "Cauchy_Integral_Theorem" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
9 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
10 |
begin |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
11 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
12 |
subsection\<open>Cauchy's inequality and more versions of Liouville\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
13 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
14 |
lemma Cauchy_higher_deriv_bound: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
15 |
assumes holf: "f holomorphic_on (ball z r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
16 |
and contf: "continuous_on (cball z r) f" |
65036
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
paulson <lp15@cam.ac.uk>
parents:
64394
diff
changeset
|
17 |
and fin : "\<And>w. w \<in> ball z r \<Longrightarrow> f w \<in> ball y B0" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
18 |
and "0 < r" and "0 < n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
19 |
shows "norm ((deriv ^^ n) f z) \<le> (fact n) * B0 / r^n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
20 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
21 |
have "0 < B0" using \<open>0 < r\<close> fin [of z] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
22 |
by (metis ball_eq_empty ex_in_conv fin not_less) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
23 |
have le_B0: "\<And>w. cmod (w - z) \<le> r \<Longrightarrow> cmod (f w - y) \<le> B0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
24 |
apply (rule continuous_on_closure_norm_le [of "ball z r" "\<lambda>w. f w - y"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
25 |
apply (auto simp: \<open>0 < r\<close> dist_norm norm_minus_commute) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
26 |
apply (rule continuous_intros contf)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
27 |
using fin apply (simp add: dist_commute dist_norm less_eq_real_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
28 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
29 |
have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w) z - (deriv ^^ n) (\<lambda>w. y) z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
30 |
using \<open>0 < n\<close> by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
31 |
also have "... = (deriv ^^ n) (\<lambda>w. f w - y) z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
32 |
by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \<open>0 < r\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
33 |
finally have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w - y) z" . |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
34 |
have contf': "continuous_on (cball z r) (\<lambda>u. f u - y)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
35 |
by (rule contf continuous_intros)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
36 |
have holf': "(\<lambda>u. (f u - y)) holomorphic_on (ball z r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
37 |
by (simp add: holf holomorphic_on_diff) |
63040 | 38 |
define a where "a = (2 * pi)/(fact n)" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
39 |
have "0 < a" by (simp add: a_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
40 |
have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
41 |
using \<open>0 < r\<close> by (simp add: a_def divide_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
42 |
have der_dif: "(deriv ^^ n) (\<lambda>w. f w - y) z = (deriv ^^ n) f z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
43 |
using \<open>0 < r\<close> \<open>0 < n\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
44 |
by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) |
63589 | 45 |
have "norm ((2 * of_real pi * \<i>)/(fact n) * (deriv ^^ n) (\<lambda>w. f w - y) z) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
46 |
\<le> (B0/r^(Suc n)) * (2 * pi * r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
47 |
apply (rule has_contour_integral_bound_circlepath [of "(\<lambda>u. (f u - y)/(u - z)^(Suc n))" _ z]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
48 |
using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
49 |
using \<open>0 < B0\<close> \<open>0 < r\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
50 |
apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
51 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
52 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
53 |
using \<open>0 < r\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
54 |
by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
55 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
56 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
57 |
proposition Cauchy_inequality: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
58 |
assumes holf: "f holomorphic_on (ball \<xi> r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
59 |
and contf: "continuous_on (cball \<xi> r) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
60 |
and "0 < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
61 |
and nof: "\<And>x. norm(\<xi>-x) = r \<Longrightarrow> norm(f x) \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
62 |
shows "norm ((deriv ^^ n) f \<xi>) \<le> (fact n) * B / r^n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
63 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
64 |
obtain x where "norm (\<xi>-x) = r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
65 |
by (metis abs_of_nonneg add_diff_cancel_left' \<open>0 < r\<close> diff_add_cancel |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
66 |
dual_order.strict_implies_order norm_of_real) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
67 |
then have "0 \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
68 |
by (metis nof norm_not_less_zero not_le order_trans) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
69 |
have "((\<lambda>u. f u / (u - \<xi>) ^ Suc n) has_contour_integral (2 * pi) * \<i> / fact n * (deriv ^^ n) f \<xi>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
70 |
(circlepath \<xi> r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
71 |
apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
72 |
using \<open>0 < r\<close> by simp |
63589 | 73 |
then have "norm ((2 * pi * \<i>)/(fact n) * (deriv ^^ n) f \<xi>) \<le> (B / r^(Suc n)) * (2 * pi * r)" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
74 |
apply (rule has_contour_integral_bound_circlepath) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
75 |
using \<open>0 \<le> B\<close> \<open>0 < r\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
76 |
apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
77 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
78 |
then show ?thesis using \<open>0 < r\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
79 |
by (simp add: norm_divide norm_mult field_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
80 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
81 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
82 |
proposition Liouville_polynomial: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
83 |
assumes holf: "f holomorphic_on UNIV" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
84 |
and nof: "\<And>z. A \<le> norm z \<Longrightarrow> norm(f z) \<le> B * norm z ^ n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
85 |
shows "f \<xi> = (\<Sum>k\<le>n. (deriv^^k) f 0 / fact k * \<xi> ^ k)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
86 |
proof (cases rule: le_less_linear [THEN disjE]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
87 |
assume "B \<le> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
88 |
then have "\<And>z. A \<le> norm z \<Longrightarrow> norm(f z) = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
89 |
by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
90 |
then have f0: "(f \<longlongrightarrow> 0) at_infinity" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
91 |
using Lim_at_infinity by force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
92 |
then have [simp]: "f = (\<lambda>w. 0)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
93 |
using Liouville_weak [OF holf, of 0] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
94 |
by (simp add: eventually_at_infinity f0) meson |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
95 |
show ?thesis by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
96 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
97 |
assume "0 < B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
98 |
have "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * (\<xi> - 0)^k) sums f \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
99 |
apply (rule holomorphic_power_series [where r = "norm \<xi> + 1"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
100 |
using holf holomorphic_on_subset apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
101 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
102 |
then have sumsf: "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * \<xi>^k) sums f \<xi>)" by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
103 |
have "(deriv ^^ k) f 0 / fact k * \<xi> ^ k = 0" if "k>n" for k |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
104 |
proof (cases "(deriv ^^ k) f 0 = 0") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
105 |
case True then show ?thesis by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
106 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
107 |
case False |
63040 | 108 |
define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))" |
109 |
have "1 \<le> abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))" |
|
110 |
using \<open>0 < B\<close> by simp |
|
111 |
then have wge1: "1 \<le> norm w" |
|
112 |
by (metis norm_of_real w_def) |
|
113 |
then have "w \<noteq> 0" by auto |
|
114 |
have kB: "0 < fact k * B" |
|
115 |
using \<open>0 < B\<close> by simp |
|
116 |
then have "0 \<le> fact k * B / cmod ((deriv ^^ k) f 0)" |
|
117 |
by simp |
|
118 |
then have wgeA: "A \<le> cmod w" |
|
119 |
by (simp only: w_def norm_of_real) |
|
120 |
have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))" |
|
121 |
using \<open>0 < B\<close> by simp |
|
122 |
then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" |
|
123 |
by (metis norm_of_real w_def) |
|
124 |
then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" |
|
125 |
using False by (simp add: divide_simps mult.commute split: if_split_asm) |
|
126 |
also have "... \<le> fact k * (B * norm w ^ n) / norm w ^ k" |
|
127 |
apply (rule Cauchy_inequality) |
|
128 |
using holf holomorphic_on_subset apply force |
|
129 |
using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast |
|
130 |
using \<open>w \<noteq> 0\<close> apply (simp add:) |
|
131 |
by (metis nof wgeA dist_0_norm dist_norm) |
|
132 |
also have "... = fact k * (B * 1 / cmod w ^ (k-n))" |
|
133 |
apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) |
|
134 |
using \<open>k>n\<close> \<open>w \<noteq> 0\<close> \<open>0 < B\<close> apply (simp add: divide_simps semiring_normalization_rules) |
|
135 |
done |
|
136 |
also have "... = fact k * B / cmod w ^ (k-n)" |
|
137 |
by simp |
|
138 |
finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . |
|
139 |
then have "1 / cmod w < 1 / cmod w ^ (k - n)" |
|
140 |
by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) |
|
141 |
then have "cmod w ^ (k - n) < cmod w" |
|
142 |
by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) |
|
143 |
with self_le_power [OF wge1] have False |
|
144 |
by (meson diff_is_0_eq not_gr0 not_le that) |
|
145 |
then show ?thesis by blast |
|
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
146 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
147 |
then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \<xi> ^ (k + Suc n) = 0" for k |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
148 |
using not_less_eq by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
149 |
then have "(\<lambda>i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \<xi> ^ (i + Suc n)) sums 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
150 |
by (rule sums_0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
151 |
with sums_split_initial_segment [OF sumsf, where n = "Suc n"] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
152 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
153 |
using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
154 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
155 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
156 |
text\<open>Every bounded entire function is a constant function.\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
157 |
theorem Liouville_theorem: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
158 |
assumes holf: "f holomorphic_on UNIV" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
159 |
and bf: "bounded (range f)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
160 |
obtains c where "\<And>z. f z = c" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
161 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
162 |
obtain B where "\<And>z. cmod (f z) \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
163 |
by (meson bf bounded_pos rangeI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
164 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
165 |
using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
166 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
167 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
168 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
169 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
170 |
text\<open>A holomorphic function f has only isolated zeros unless f is 0.\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
171 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
172 |
proposition powser_0_nonzero: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
173 |
fixes a :: "nat \<Rightarrow> 'a::{real_normed_field,banach}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
174 |
assumes r: "0 < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
175 |
and sm: "\<And>x. norm (x - \<xi>) < r \<Longrightarrow> (\<lambda>n. a n * (x - \<xi>) ^ n) sums (f x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
176 |
and [simp]: "f \<xi> = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
177 |
and m0: "a m \<noteq> 0" and "m>0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
178 |
obtains s where "0 < s" and "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
179 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
180 |
have "r \<le> conv_radius a" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
181 |
using sm sums_summable by (auto simp: le_conv_radius_iff [where \<xi>=\<xi>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
182 |
obtain m where am: "a m \<noteq> 0" and az [simp]: "(\<And>n. n<m \<Longrightarrow> a n = 0)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
183 |
apply (rule_tac m = "LEAST n. a n \<noteq> 0" in that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
184 |
using m0 |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
185 |
apply (rule LeastI2) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
186 |
apply (fastforce intro: dest!: not_less_Least)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
187 |
done |
63040 | 188 |
define b where "b i = a (i+m) / a m" for i |
189 |
define g where "g x = suminf (\<lambda>i. b i * (x - \<xi>) ^ i)" for x |
|
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
190 |
have [simp]: "b 0 = 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
191 |
by (simp add: am b_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
192 |
{ fix x::'a |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
193 |
assume "norm (x - \<xi>) < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
194 |
then have "(\<lambda>n. (a m * (x - \<xi>)^m) * (b n * (x - \<xi>)^n)) sums (f x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
195 |
using am az sm sums_zero_iff_shift [of m "(\<lambda>n. a n * (x - \<xi>) ^ n)" "f x"] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
196 |
by (simp add: b_def monoid_mult_class.power_add algebra_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
197 |
then have "x \<noteq> \<xi> \<Longrightarrow> (\<lambda>n. b n * (x - \<xi>)^n) sums (f x / (a m * (x - \<xi>)^m))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
198 |
using am by (simp add: sums_mult_D) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
199 |
} note bsums = this |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
200 |
then have "norm (x - \<xi>) < r \<Longrightarrow> summable (\<lambda>n. b n * (x - \<xi>)^n)" for x |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
201 |
using sums_summable by (cases "x=\<xi>") auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
202 |
then have "r \<le> conv_radius b" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
203 |
by (simp add: le_conv_radius_iff [where \<xi>=\<xi>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
204 |
then have "r/2 < conv_radius b" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
205 |
using not_le order_trans r by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
206 |
then have "continuous_on (cball \<xi> (r/2)) g" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
207 |
using powser_continuous_suminf [of "r/2" b \<xi>] by (simp add: g_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
208 |
then obtain s where "s>0" "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> dist (g x) (g \<xi>) < 1/2" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
209 |
apply (rule continuous_onE [where x=\<xi> and e = "1/2"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
210 |
using r apply (auto simp: norm_minus_commute dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
211 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
212 |
moreover have "g \<xi> = 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
213 |
by (simp add: g_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
214 |
ultimately have gnz: "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> (g x) \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
215 |
by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
216 |
have "f x \<noteq> 0" if "x \<noteq> \<xi>" "norm (x - \<xi>) \<le> s" "norm (x - \<xi>) \<le> r/2" for x |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
217 |
using bsums [of x] that gnz [of x] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
218 |
apply (auto simp: g_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
219 |
using r sums_iff by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
220 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
221 |
apply (rule_tac s="min s (r/2)" in that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
222 |
using \<open>0 < r\<close> \<open>0 < s\<close> by (auto simp: dist_commute dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
223 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
224 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
225 |
proposition isolated_zeros: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
226 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
227 |
and "open S" "connected S" "\<xi> \<in> S" "f \<xi> = 0" "\<beta> \<in> S" "f \<beta> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
228 |
obtains r where "0 < r" "ball \<xi> r \<subseteq> S" "\<And>z. z \<in> ball \<xi> r - {\<xi>} \<Longrightarrow> f z \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
229 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
230 |
obtain r where "0 < r" and r: "ball \<xi> r \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
231 |
using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_ball_eq by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
232 |
have powf: "((\<lambda>n. (deriv ^^ n) f \<xi> / (fact n) * (z - \<xi>)^n) sums f z)" if "z \<in> ball \<xi> r" for z |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
233 |
apply (rule holomorphic_power_series [OF _ that]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
234 |
apply (rule holomorphic_on_subset [OF holf r]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
235 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
236 |
obtain m where m: "(deriv ^^ m) f \<xi> / (fact m) \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
237 |
using holomorphic_fun_eq_0_on_connected [OF holf \<open>open S\<close> \<open>connected S\<close> _ \<open>\<xi> \<in> S\<close> \<open>\<beta> \<in> S\<close>] \<open>f \<beta> \<noteq> 0\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
238 |
by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
239 |
then have "m \<noteq> 0" using assms(5) funpow_0 by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
240 |
obtain s where "0 < s" and s: "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
241 |
apply (rule powser_0_nonzero [OF \<open>0 < r\<close> powf \<open>f \<xi> = 0\<close> m]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
242 |
using \<open>m \<noteq> 0\<close> by (auto simp: dist_commute dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
243 |
have "0 < min r s" by (simp add: \<open>0 < r\<close> \<open>0 < s\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
244 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
245 |
apply (rule that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
246 |
using r s by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
247 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
248 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
249 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
250 |
proposition analytic_continuation: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
251 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
252 |
and S: "open S" "connected S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
253 |
and "U \<subseteq> S" "\<xi> \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
254 |
and "\<xi> islimpt U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
255 |
and fU0 [simp]: "\<And>z. z \<in> U \<Longrightarrow> f z = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
256 |
and "w \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
257 |
shows "f w = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
258 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
259 |
obtain e where "0 < e" and e: "cball \<xi> e \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
260 |
using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_cball_eq by blast |
63040 | 261 |
define T where "T = cball \<xi> e \<inter> U" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
262 |
have contf: "continuous_on (closure T) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
263 |
by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
264 |
holomorphic_on_subset inf.cobounded1) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
265 |
have fT0 [simp]: "\<And>x. x \<in> T \<Longrightarrow> f x = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
266 |
by (simp add: T_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
267 |
have "\<And>r. \<lbrakk>\<forall>e>0. \<exists>x'\<in>U. x' \<noteq> \<xi> \<and> dist x' \<xi> < e; 0 < r\<rbrakk> \<Longrightarrow> \<exists>x'\<in>cball \<xi> e \<inter> U. x' \<noteq> \<xi> \<and> dist x' \<xi> < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
268 |
by (metis \<open>0 < e\<close> IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
269 |
then have "\<xi> islimpt T" using \<open>\<xi> islimpt U\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
270 |
by (auto simp: T_def islimpt_approachable) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
271 |
then have "\<xi> \<in> closure T" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
272 |
by (simp add: closure_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
273 |
then have "f \<xi> = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
274 |
by (auto simp: continuous_constant_on_closure [OF contf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
275 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
276 |
apply (rule ccontr) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
277 |
apply (rule isolated_zeros [OF holf \<open>open S\<close> \<open>connected S\<close> \<open>\<xi> \<in> S\<close> \<open>f \<xi> = 0\<close> \<open>w \<in> S\<close>], assumption) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
278 |
by (metis open_ball \<open>\<xi> islimpt T\<close> centre_in_ball fT0 insertE insert_Diff islimptE) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
279 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
280 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
281 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
282 |
subsection\<open>Open mapping theorem\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
283 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
284 |
lemma holomorphic_contract_to_zero: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
285 |
assumes contf: "continuous_on (cball \<xi> r) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
286 |
and holf: "f holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
287 |
and "0 < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
288 |
and norm_less: "\<And>z. norm(\<xi> - z) = r \<Longrightarrow> norm(f \<xi>) < norm(f z)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
289 |
obtains z where "z \<in> ball \<xi> r" "f z = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
290 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
291 |
{ assume fnz: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
292 |
then have "0 < norm (f \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
293 |
by (simp add: \<open>0 < r\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
294 |
have fnz': "\<And>w. w \<in> cball \<xi> r \<Longrightarrow> f w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
295 |
by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
296 |
have "frontier(cball \<xi> r) \<noteq> {}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
297 |
using \<open>0 < r\<close> by simp |
63040 | 298 |
define g where [abs_def]: "g z = inverse (f z)" for z |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
299 |
have contg: "continuous_on (cball \<xi> r) g" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
300 |
unfolding g_def using contf continuous_on_inverse fnz' by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
301 |
have holg: "g holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
302 |
unfolding g_def using fnz holf holomorphic_on_inverse by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
303 |
have "frontier (cball \<xi> r) \<subseteq> cball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
304 |
by (simp add: subset_iff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
305 |
then have contf': "continuous_on (frontier (cball \<xi> r)) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
306 |
and contg': "continuous_on (frontier (cball \<xi> r)) g" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
307 |
by (blast intro: contf contg continuous_on_subset)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
308 |
have froc: "frontier(cball \<xi> r) \<noteq> {}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
309 |
using \<open>0 < r\<close> by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
310 |
moreover have "continuous_on (frontier (cball \<xi> r)) (norm o f)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
311 |
using contf' continuous_on_compose continuous_on_norm_id by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
312 |
ultimately obtain w where w: "w \<in> frontier(cball \<xi> r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
313 |
and now: "\<And>x. x \<in> frontier(cball \<xi> r) \<Longrightarrow> norm (f w) \<le> norm (f x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
314 |
apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
315 |
apply (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
316 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
317 |
then have fw: "0 < norm (f w)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
318 |
by (simp add: fnz') |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
319 |
have "continuous_on (frontier (cball \<xi> r)) (norm o g)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
320 |
using contg' continuous_on_compose continuous_on_norm_id by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
321 |
then obtain v where v: "v \<in> frontier(cball \<xi> r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
322 |
and nov: "\<And>x. x \<in> frontier(cball \<xi> r) \<Longrightarrow> norm (g v) \<ge> norm (g x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
323 |
apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
324 |
apply (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
325 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
326 |
then have fv: "0 < norm (f v)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
327 |
by (simp add: fnz') |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
328 |
have "norm ((deriv ^^ 0) g \<xi>) \<le> fact 0 * norm (g v) / r ^ 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
329 |
by (rule Cauchy_inequality [OF holg contg \<open>0 < r\<close>]) (simp add: dist_norm nov) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
330 |
then have "cmod (g \<xi>) \<le> norm (g v)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
331 |
by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
332 |
with w have wr: "norm (\<xi> - w) = r" and nfw: "norm (f w) \<le> norm (f \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
333 |
apply (simp_all add: dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
334 |
by (metis \<open>0 < cmod (f \<xi>)\<close> g_def less_imp_inverse_less norm_inverse not_le now order_trans v) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
335 |
with fw have False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
336 |
using norm_less by force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
337 |
} |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
338 |
with that show ?thesis by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
339 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
340 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
341 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
342 |
theorem open_mapping_thm: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
343 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
344 |
and S: "open S" "connected S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
345 |
and "open U" "U \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
346 |
and fne: "~ f constant_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
347 |
shows "open (f ` U)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
348 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
349 |
have *: "open (f ` U)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
350 |
if "U \<noteq> {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\<And>x. \<exists>y \<in> U. f y \<noteq> x" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
351 |
for U |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
352 |
proof (clarsimp simp: open_contains_ball) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
353 |
fix \<xi> assume \<xi>: "\<xi> \<in> U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
354 |
show "\<exists>e>0. ball (f \<xi>) e \<subseteq> f ` U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
355 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
356 |
have hol: "(\<lambda>z. f z - f \<xi>) holomorphic_on U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
357 |
by (rule holomorphic_intros that)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
358 |
obtain s where "0 < s" and sbU: "ball \<xi> s \<subseteq> U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
359 |
and sne: "\<And>z. z \<in> ball \<xi> s - {\<xi>} \<Longrightarrow> (\<lambda>z. f z - f \<xi>) z \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
360 |
using isolated_zeros [OF hol U \<xi>] by (metis fneU right_minus_eq) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
361 |
obtain r where "0 < r" and r: "cball \<xi> r \<subseteq> ball \<xi> s" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
362 |
apply (rule_tac r="s/2" in that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
363 |
using \<open>0 < s\<close> by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
364 |
have "cball \<xi> r \<subseteq> U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
365 |
using sbU r by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
366 |
then have frsbU: "frontier (cball \<xi> r) \<subseteq> U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
367 |
using Diff_subset frontier_def order_trans by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
368 |
then have cof: "compact (frontier(cball \<xi> r))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
369 |
by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
370 |
have frne: "frontier (cball \<xi> r) \<noteq> {}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
371 |
using \<open>0 < r\<close> by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
372 |
have contfr: "continuous_on (frontier (cball \<xi> r)) (\<lambda>z. norm (f z - f \<xi>))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
373 |
apply (rule continuous_on_compose2 [OF Complex_Analysis_Basics.continuous_on_norm_id]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
374 |
using hol frsbU holomorphic_on_imp_continuous_on holomorphic_on_subset by blast+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
375 |
obtain w where "norm (\<xi> - w) = r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
376 |
and w: "(\<And>z. norm (\<xi> - z) = r \<Longrightarrow> norm (f w - f \<xi>) \<le> norm(f z - f \<xi>))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
377 |
apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
378 |
apply (simp add: dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
379 |
done |
63040 | 380 |
moreover define \<epsilon> where "\<epsilon> \<equiv> norm (f w - f \<xi>) / 3" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
381 |
ultimately have "0 < \<epsilon>" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
382 |
using \<open>0 < r\<close> dist_complex_def r sne by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
383 |
have "ball (f \<xi>) \<epsilon> \<subseteq> f ` U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
384 |
proof |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
385 |
fix \<gamma> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
386 |
assume \<gamma>: "\<gamma> \<in> ball (f \<xi>) \<epsilon>" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
387 |
have *: "cmod (\<gamma> - f \<xi>) < cmod (\<gamma> - f z)" if "cmod (\<xi> - z) = r" for z |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
388 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
389 |
have lt: "cmod (f w - f \<xi>) / 3 < cmod (\<gamma> - f z)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
390 |
using w [OF that] \<gamma> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
391 |
using dist_triangle2 [of "f \<xi>" "\<gamma>" "f z"] dist_triangle2 [of "f \<xi>" "f z" \<gamma>] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
392 |
by (simp add: \<epsilon>_def dist_norm norm_minus_commute) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
393 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
394 |
by (metis \<epsilon>_def dist_commute dist_norm less_trans lt mem_ball \<gamma>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
395 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
396 |
have "continuous_on (cball \<xi> r) (\<lambda>z. \<gamma> - f z)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
397 |
apply (rule continuous_intros)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
398 |
using \<open>cball \<xi> r \<subseteq> U\<close> \<open>f holomorphic_on U\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
399 |
apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
400 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
401 |
moreover have "(\<lambda>z. \<gamma> - f z) holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
402 |
apply (rule holomorphic_intros)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
403 |
apply (metis \<open>cball \<xi> r \<subseteq> U\<close> \<open>f holomorphic_on U\<close> holomorphic_on_subset interior_cball interior_subset) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
404 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
405 |
ultimately obtain z where "z \<in> ball \<xi> r" "\<gamma> - f z = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
406 |
apply (rule holomorphic_contract_to_zero) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
407 |
apply (blast intro!: \<open>0 < r\<close> *)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
408 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
409 |
then show "\<gamma> \<in> f ` U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
410 |
using \<open>cball \<xi> r \<subseteq> U\<close> by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
411 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
412 |
then show ?thesis using \<open>0 < \<epsilon>\<close> by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
413 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
414 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
415 |
have "open (f ` X)" if "X \<in> components U" for X |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
416 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
417 |
have holfU: "f holomorphic_on U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
418 |
using \<open>U \<subseteq> S\<close> holf holomorphic_on_subset by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
419 |
have "X \<noteq> {}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
420 |
using that by (simp add: in_components_nonempty) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
421 |
moreover have "open X" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
422 |
using that \<open>open U\<close> open_components by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
423 |
moreover have "connected X" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
424 |
using that in_components_maximal by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
425 |
moreover have "f holomorphic_on X" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
426 |
by (meson that holfU holomorphic_on_subset in_components_maximal) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
427 |
moreover have "\<exists>y\<in>X. f y \<noteq> x" for x |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
428 |
proof (rule ccontr) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
429 |
assume not: "\<not> (\<exists>y\<in>X. f y \<noteq> x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
430 |
have "X \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
431 |
using \<open>U \<subseteq> S\<close> in_components_subset that by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
432 |
obtain w where w: "w \<in> X" using \<open>X \<noteq> {}\<close> by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
433 |
have wis: "w islimpt X" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
434 |
using w \<open>open X\<close> interior_eq by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
435 |
have hol: "(\<lambda>z. f z - x) holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
436 |
by (simp add: holf holomorphic_on_diff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
437 |
with fne [unfolded constant_on_def] analytic_continuation [OF hol S \<open>X \<subseteq> S\<close> _ wis] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
438 |
not \<open>X \<subseteq> S\<close> w |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
439 |
show False by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
440 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
441 |
ultimately show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
442 |
by (rule *) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
443 |
qed |
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62837
diff
changeset
|
444 |
then have "open (f ` \<Union>components U)" |
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62837
diff
changeset
|
445 |
by (metis (no_types, lifting) imageE image_Union open_Union) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
446 |
then show ?thesis |
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62837
diff
changeset
|
447 |
by force |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
448 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
449 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
450 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
451 |
text\<open>No need for @{term S} to be connected. But the nonconstant condition is stronger.\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
452 |
corollary open_mapping_thm2: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
453 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
454 |
and S: "open S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
455 |
and "open U" "U \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
456 |
and fnc: "\<And>X. \<lbrakk>open X; X \<subseteq> S; X \<noteq> {}\<rbrakk> \<Longrightarrow> ~ f constant_on X" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
457 |
shows "open (f ` U)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
458 |
proof - |
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62837
diff
changeset
|
459 |
have "S = \<Union>(components S)" by simp |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
460 |
with \<open>U \<subseteq> S\<close> have "U = (\<Union>C \<in> components S. C \<inter> U)" by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
461 |
then have "f ` U = (\<Union>C \<in> components S. f ` (C \<inter> U))" |
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62837
diff
changeset
|
462 |
using image_UN by fastforce |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
463 |
moreover |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
464 |
{ fix C assume "C \<in> components S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
465 |
with S \<open>C \<in> components S\<close> open_components in_components_connected |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
466 |
have C: "open C" "connected C" by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
467 |
have "C \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
468 |
by (metis \<open>C \<in> components S\<close> in_components_maximal) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
469 |
have nf: "\<not> f constant_on C" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
470 |
apply (rule fnc) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
471 |
using C \<open>C \<subseteq> S\<close> \<open>C \<in> components S\<close> in_components_nonempty by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
472 |
have "f holomorphic_on C" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
473 |
by (metis holf holomorphic_on_subset \<open>C \<subseteq> S\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
474 |
then have "open (f ` (C \<inter> U))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
475 |
apply (rule open_mapping_thm [OF _ C _ _ nf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
476 |
apply (simp add: C \<open>open U\<close> open_Int, blast) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
477 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
478 |
} ultimately show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
479 |
by force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
480 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
481 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
482 |
corollary open_mapping_thm3: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
483 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
484 |
and "open S" and injf: "inj_on f S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
485 |
shows "open (f ` S)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
486 |
apply (rule open_mapping_thm2 [OF holf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
487 |
using assms |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
488 |
apply (simp_all add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
489 |
using injective_not_constant subset_inj_on by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
490 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
491 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
492 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
493 |
subsection\<open>Maximum Modulus Principle\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
494 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
495 |
text\<open>If @{term f} is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
496 |
properly within the domain of @{term f}.\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
497 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
498 |
proposition maximum_modulus_principle: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
499 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
500 |
and S: "open S" "connected S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
501 |
and "open U" "U \<subseteq> S" "\<xi> \<in> U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
502 |
and no: "\<And>z. z \<in> U \<Longrightarrow> norm(f z) \<le> norm(f \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
503 |
shows "f constant_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
504 |
proof (rule ccontr) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
505 |
assume "\<not> f constant_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
506 |
then have "open (f ` U)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
507 |
using open_mapping_thm assms by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
508 |
moreover have "~ open (f ` U)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
509 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
510 |
have "\<exists>t. cmod (f \<xi> - t) < e \<and> t \<notin> f ` U" if "0 < e" for e |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
511 |
apply (rule_tac x="if 0 < Re(f \<xi>) then f \<xi> + (e/2) else f \<xi> - (e/2)" in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
512 |
using that |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
513 |
apply (simp add: dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
514 |
apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
515 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
516 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
517 |
unfolding open_contains_ball by (metis \<open>\<xi> \<in> U\<close> contra_subsetD dist_norm imageI mem_ball) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
518 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
519 |
ultimately show False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
520 |
by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
521 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
522 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
523 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
524 |
proposition maximum_modulus_frontier: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
525 |
assumes holf: "f holomorphic_on (interior S)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
526 |
and contf: "continuous_on (closure S) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
527 |
and bos: "bounded S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
528 |
and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> norm(f z) \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
529 |
and "\<xi> \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
530 |
shows "norm(f \<xi>) \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
531 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
532 |
have "compact (closure S)" using bos |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
533 |
by (simp add: bounded_closure compact_eq_bounded_closed) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
534 |
moreover have "continuous_on (closure S) (cmod \<circ> f)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
535 |
using contf continuous_on_compose continuous_on_norm_id by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
536 |
ultimately obtain z where zin: "z \<in> closure S" and z: "\<And>y. y \<in> closure S \<Longrightarrow> (cmod \<circ> f) y \<le> (cmod \<circ> f) z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
537 |
using continuous_attains_sup [of "closure S" "norm o f"] \<open>\<xi> \<in> S\<close> by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
538 |
then consider "z \<in> frontier S" | "z \<in> interior S" using frontier_def by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
539 |
then have "norm(f z) \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
540 |
proof cases |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
541 |
case 1 then show ?thesis using leB by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
542 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
543 |
case 2 |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
544 |
have zin: "z \<in> connected_component_set (interior S) z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
545 |
by (simp add: 2) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
546 |
have "f constant_on (connected_component_set (interior S) z)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
547 |
apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
548 |
apply (metis connected_component_subset holf holomorphic_on_subset) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
549 |
apply (simp_all add: open_connected_component) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
550 |
by (metis closure_subset comp_eq_dest_lhs interior_subset subsetCE z connected_component_in) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
551 |
then obtain c where c: "\<And>w. w \<in> connected_component_set (interior S) z \<Longrightarrow> f w = c" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
552 |
by (auto simp: constant_on_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
553 |
have "f ` closure(connected_component_set (interior S) z) \<subseteq> {c}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
554 |
apply (rule image_closure_subset) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
555 |
apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
556 |
using c |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
557 |
apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
558 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
559 |
then have cc: "\<And>w. w \<in> closure(connected_component_set (interior S) z) \<Longrightarrow> f w = c" by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
560 |
have "frontier(connected_component_set (interior S) z) \<noteq> {}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
561 |
apply (simp add: frontier_eq_empty) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
562 |
by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
563 |
then obtain w where w: "w \<in> frontier(connected_component_set (interior S) z)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
564 |
by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
565 |
then have "norm (f z) = norm (f w)" by (simp add: "2" c cc frontier_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
566 |
also have "... \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
567 |
apply (rule leB) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
568 |
using w |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
569 |
using frontier_interior_subset frontier_of_connected_component_subset by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
570 |
finally show ?thesis . |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
571 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
572 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
573 |
using z \<open>\<xi> \<in> S\<close> closure_subset by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
574 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
575 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
576 |
corollary maximum_real_frontier: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
577 |
assumes holf: "f holomorphic_on (interior S)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
578 |
and contf: "continuous_on (closure S) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
579 |
and bos: "bounded S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
580 |
and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> Re(f z) \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
581 |
and "\<xi> \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
582 |
shows "Re(f \<xi>) \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
583 |
using maximum_modulus_frontier [of "exp o f" S "exp B"] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
584 |
Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
585 |
by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
586 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
587 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
588 |
subsection\<open>Factoring out a zero according to its order\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
589 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
590 |
lemma holomorphic_factor_order_of_zero: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
591 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
592 |
and os: "open S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
593 |
and "\<xi> \<in> S" "0 < n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
594 |
and dnz: "(deriv ^^ n) f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
595 |
and dfz: "\<And>i. \<lbrakk>0 < i; i < n\<rbrakk> \<Longrightarrow> (deriv ^^ i) f \<xi> = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
596 |
obtains g r where "0 < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
597 |
"g holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
598 |
"\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>)^n * g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
599 |
"\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
600 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
601 |
obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
602 |
then have holfb: "f holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
603 |
using holf holomorphic_on_subset by blast |
63040 | 604 |
define g where "g w = suminf (\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i)" for w |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
605 |
have sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
606 |
and feq: "f w - f \<xi> = (w - \<xi>)^n * g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
607 |
if w: "w \<in> ball \<xi> r" for w |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
608 |
proof - |
63040 | 609 |
define powf where "powf = (\<lambda>i. (deriv ^^ i) f \<xi>/(fact i) * (w - \<xi>)^i)" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
610 |
have sing: "{..<n} - {i. powf i = 0} = (if f \<xi> = 0 then {} else {0})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
611 |
unfolding powf_def using \<open>0 < n\<close> dfz by (auto simp: dfz; metis funpow_0 not_gr0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
612 |
have "powf sums f w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
613 |
unfolding powf_def by (rule holomorphic_power_series [OF holfb w]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
614 |
moreover have "(\<Sum>i<n. powf i) = f \<xi>" |
64267 | 615 |
apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric]) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
616 |
apply (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
617 |
apply (simp only: dfz sing) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
618 |
apply (simp add: powf_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
619 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
620 |
ultimately have fsums: "(\<lambda>i. powf (i+n)) sums (f w - f \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
621 |
using w sums_iff_shift' by metis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
622 |
then have *: "summable (\<lambda>i. (w - \<xi>) ^ n * ((deriv ^^ (i + n)) f \<xi> * (w - \<xi>) ^ i / fact (i + n)))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
623 |
unfolding powf_def using sums_summable |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
624 |
by (auto simp: power_add mult_ac) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
625 |
have "summable (\<lambda>i. (deriv ^^ (i + n)) f \<xi> * (w - \<xi>) ^ i / fact (i + n))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
626 |
proof (cases "w=\<xi>") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
627 |
case False then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
628 |
using summable_mult [OF *, of "1 / (w - \<xi>) ^ n"] by (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
629 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
630 |
case True then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
631 |
by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
632 |
split: if_split_asm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
633 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
634 |
then show sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
635 |
by (simp add: summable_sums_iff g_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
636 |
show "f w - f \<xi> = (w - \<xi>)^n * g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
637 |
apply (rule sums_unique2) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
638 |
apply (rule fsums [unfolded powf_def]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
639 |
using sums_mult [OF sumsg, of "(w - \<xi>) ^ n"] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
640 |
by (auto simp: power_add mult_ac) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
641 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
642 |
then have holg: "g holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
643 |
by (meson sumsg power_series_holomorphic) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
644 |
then have contg: "continuous_on (ball \<xi> r) g" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
645 |
by (blast intro: holomorphic_on_imp_continuous_on) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
646 |
have "g \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
647 |
using dnz unfolding g_def |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
648 |
by (subst suminf_finite [of "{0}"]) auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
649 |
obtain d where "0 < d" and d: "\<And>w. w \<in> ball \<xi> d \<Longrightarrow> g w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
650 |
apply (rule exE [OF continuous_on_avoid [OF contg _ \<open>g \<xi> \<noteq> 0\<close>]]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
651 |
using \<open>0 < r\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
652 |
apply force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
653 |
by (metis \<open>0 < r\<close> less_trans mem_ball not_less_iff_gr_or_eq) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
654 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
655 |
apply (rule that [where g=g and r ="min r d"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
656 |
using \<open>0 < r\<close> \<open>0 < d\<close> holg |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
657 |
apply (auto simp: feq holomorphic_on_subset subset_ball d) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
658 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
659 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
660 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
661 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
662 |
lemma holomorphic_factor_order_of_zero_strong: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
663 |
assumes holf: "f holomorphic_on S" "open S" "\<xi> \<in> S" "0 < n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
664 |
and "(deriv ^^ n) f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
665 |
and "\<And>i. \<lbrakk>0 < i; i < n\<rbrakk> \<Longrightarrow> (deriv ^^ i) f \<xi> = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
666 |
obtains g r where "0 < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
667 |
"g holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
668 |
"\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = ((w - \<xi>) * g w) ^ n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
669 |
"\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
670 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
671 |
obtain g r where "0 < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
672 |
and holg: "g holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
673 |
and feq: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>)^n * g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
674 |
and gne: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
675 |
by (auto intro: holomorphic_factor_order_of_zero [OF assms]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
676 |
have con: "continuous_on (ball \<xi> r) (\<lambda>z. deriv g z / g z)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
677 |
by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on) |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
678 |
have cd: "\<And>x. dist \<xi> x < r \<Longrightarrow> (\<lambda>z. deriv g z / g z) field_differentiable at x" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
679 |
apply (rule derivative_intros)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
680 |
using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
681 |
apply (metis Topology_Euclidean_Space.open_ball at_within_open holg holomorphic_on_def mem_ball) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
682 |
using gne mem_ball by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
683 |
obtain h where h: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> (h has_field_derivative deriv g x / g x) (at x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
684 |
apply (rule exE [OF holomorphic_convex_primitive [of "ball \<xi> r" "{}" "\<lambda>z. deriv g z / g z"]]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
685 |
apply (auto simp: con cd) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
686 |
apply (metis open_ball at_within_open mem_ball) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
687 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
688 |
then have "continuous_on (ball \<xi> r) h" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
689 |
by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
690 |
then have con: "continuous_on (ball \<xi> r) (\<lambda>x. exp (h x) / g x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
691 |
by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
692 |
have 0: "dist \<xi> x < r \<Longrightarrow> ((\<lambda>x. exp (h x) / g x) has_field_derivative 0) (at x)" for x |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
693 |
apply (rule h derivative_eq_intros | simp)+ |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
694 |
apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2]) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
695 |
using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
696 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
697 |
obtain c where c: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> exp (h x) / g x = c" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
698 |
by (rule DERIV_zero_connected_constant [of "ball \<xi> r" "{}" "\<lambda>x. exp(h x) / g x"]) (auto simp: con 0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
699 |
have hol: "(\<lambda>z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
700 |
apply (rule holomorphic_on_compose [unfolded o_def, where g = exp]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
701 |
apply (rule holomorphic_intros)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
702 |
using h holomorphic_on_open apply blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
703 |
apply (rule holomorphic_intros)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
704 |
using \<open>0 < n\<close> apply (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
705 |
apply (rule holomorphic_intros)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
706 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
707 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
708 |
apply (rule that [where g="\<lambda>z. exp((Ln(inverse c) + h z)/n)" and r =r]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
709 |
using \<open>0 < r\<close> \<open>0 < n\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
710 |
apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
711 |
apply (rule hol) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
712 |
apply (simp add: Transcendental.exp_add gne) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
713 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
714 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
715 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
716 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
717 |
lemma |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
718 |
fixes k :: "'a::wellorder" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
719 |
assumes a_def: "a == LEAST x. P x" and P: "P k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
720 |
shows def_LeastI: "P a" and def_Least_le: "a \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
721 |
unfolding a_def |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
722 |
by (rule LeastI Least_le; rule P)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
723 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
724 |
lemma holomorphic_factor_zero_nonconstant: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
725 |
assumes holf: "f holomorphic_on S" and S: "open S" "connected S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
726 |
and "\<xi> \<in> S" "f \<xi> = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
727 |
and nonconst: "\<And>c. \<exists>z \<in> S. f z \<noteq> c" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
728 |
obtains g r n |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
729 |
where "0 < n" "0 < r" "ball \<xi> r \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
730 |
"g holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
731 |
"\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w = (w - \<xi>)^n * g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
732 |
"\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
733 |
proof (cases "\<forall>n>0. (deriv ^^ n) f \<xi> = 0") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
734 |
case True then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
735 |
using holomorphic_fun_eq_const_on_connected [OF holf S _ \<open>\<xi> \<in> S\<close>] nonconst by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
736 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
737 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
738 |
then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0" by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
739 |
obtain r0 where "r0 > 0" "ball \<xi> r0 \<subseteq> S" using S openE \<open>\<xi> \<in> S\<close> by auto |
63040 | 740 |
define n where "n \<equiv> LEAST n. (deriv ^^ n) f \<xi> \<noteq> 0" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
741 |
have n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
742 |
by (rule def_LeastI [OF n_def]) (rule n0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
743 |
then have "0 < n" using \<open>f \<xi> = 0\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
744 |
using funpow_0 by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
745 |
have n_min: "\<And>k. k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
746 |
using def_Least_le [OF n_def] not_le by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
747 |
then obtain g r1 |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
748 |
where "0 < r1" "g holomorphic_on ball \<xi> r1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
749 |
"\<And>w. w \<in> ball \<xi> r1 \<Longrightarrow> f w = (w - \<xi>) ^ n * g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
750 |
"\<And>w. w \<in> ball \<xi> r1 \<Longrightarrow> g w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
751 |
by (auto intro: holomorphic_factor_order_of_zero [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> \<open>n > 0\<close> n_ne] simp: \<open>f \<xi> = 0\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
752 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
753 |
apply (rule_tac g=g and r="min r0 r1" and n=n in that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
754 |
using \<open>0 < n\<close> \<open>0 < r0\<close> \<open>0 < r1\<close> \<open>ball \<xi> r0 \<subseteq> S\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
755 |
apply (auto simp: subset_ball intro: holomorphic_on_subset) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
756 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
757 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
758 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
759 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
760 |
lemma holomorphic_lower_bound_difference: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
761 |
assumes holf: "f holomorphic_on S" and S: "open S" "connected S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
762 |
and "\<xi> \<in> S" and "\<phi> \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
763 |
and fne: "f \<phi> \<noteq> f \<xi>" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
764 |
obtains k n r |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
765 |
where "0 < k" "0 < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
766 |
"ball \<xi> r \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
767 |
"\<And>w. w \<in> ball \<xi> r \<Longrightarrow> k * norm(w - \<xi>)^n \<le> norm(f w - f \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
768 |
proof - |
63040 | 769 |
define n where "n = (LEAST n. 0 < n \<and> (deriv ^^ n) f \<xi> \<noteq> 0)" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
770 |
obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
771 |
using fne holomorphic_fun_eq_const_on_connected [OF holf S] \<open>\<xi> \<in> S\<close> \<open>\<phi> \<in> S\<close> by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
772 |
then have "0 < n" and n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
773 |
unfolding n_def by (metis (mono_tags, lifting) LeastI)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
774 |
have n_min: "\<And>k. \<lbrakk>0 < k; k < n\<rbrakk> \<Longrightarrow> (deriv ^^ k) f \<xi> = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
775 |
unfolding n_def by (blast dest: not_less_Least) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
776 |
then obtain g r |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
777 |
where "0 < r" and holg: "g holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
778 |
and fne: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>) ^ n * g w" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
779 |
and gnz: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
780 |
by (auto intro: holomorphic_factor_order_of_zero [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> \<open>n > 0\<close> n_ne]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
781 |
obtain e where "e>0" and e: "ball \<xi> e \<subseteq> S" using assms by (blast elim!: openE) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
782 |
then have holfb: "f holomorphic_on ball \<xi> e" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
783 |
using holf holomorphic_on_subset by blast |
63040 | 784 |
define d where "d = (min e r) / 2" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
785 |
have "0 < d" using \<open>0 < r\<close> \<open>0 < e\<close> by (simp add: d_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
786 |
have "d < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
787 |
using \<open>0 < r\<close> by (auto simp: d_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
788 |
then have cbb: "cball \<xi> d \<subseteq> ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
789 |
by (auto simp: cball_subset_ball_iff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
790 |
then have "g holomorphic_on cball \<xi> d" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
791 |
by (rule holomorphic_on_subset [OF holg]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
792 |
then have "closed (g ` cball \<xi> d)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
793 |
by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
794 |
moreover have "g ` cball \<xi> d \<noteq> {}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
795 |
using \<open>0 < d\<close> by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
796 |
ultimately obtain x where x: "x \<in> g ` cball \<xi> d" and "\<And>y. y \<in> g ` cball \<xi> d \<Longrightarrow> dist 0 x \<le> dist 0 y" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
797 |
by (rule distance_attains_inf) blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
798 |
then have leg: "\<And>w. w \<in> cball \<xi> d \<Longrightarrow> norm x \<le> norm (g w)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
799 |
by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
800 |
have "ball \<xi> d \<subseteq> cball \<xi> d" by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
801 |
also have "... \<subseteq> ball \<xi> e" using \<open>0 < d\<close> d_def by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
802 |
also have "... \<subseteq> S" by (rule e) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
803 |
finally have dS: "ball \<xi> d \<subseteq> S" . |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
804 |
moreover have "x \<noteq> 0" using gnz x \<open>d < r\<close> by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
805 |
ultimately show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
806 |
apply (rule_tac k="norm x" and n=n and r=d in that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
807 |
using \<open>d < r\<close> leg |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
808 |
apply (auto simp: \<open>0 < d\<close> fne norm_mult norm_power algebra_simps mult_right_mono) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
809 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
810 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
811 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
812 |
lemma |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
813 |
assumes holf: "f holomorphic_on (S - {\<xi>})" and \<xi>: "\<xi> \<in> interior S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
814 |
shows holomorphic_on_extend_lim: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
815 |
"(\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S - {\<xi>}. g z = f z)) \<longleftrightarrow> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
816 |
((\<lambda>z. (z - \<xi>) * f z) \<longlongrightarrow> 0) (at \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
817 |
(is "?P = ?Q") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
818 |
and holomorphic_on_extend_bounded: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
819 |
"(\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S - {\<xi>}. g z = f z)) \<longleftrightarrow> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
820 |
(\<exists>B. eventually (\<lambda>z. norm(f z) \<le> B) (at \<xi>))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
821 |
(is "?P = ?R") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
822 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
823 |
obtain \<delta> where "0 < \<delta>" and \<delta>: "ball \<xi> \<delta> \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
824 |
using \<xi> mem_interior by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
825 |
have "?R" if holg: "g holomorphic_on S" and gf: "\<And>z. z \<in> S - {\<xi>} \<Longrightarrow> g z = f z" for g |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
826 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
827 |
have *: "\<forall>\<^sub>F z in at \<xi>. dist (g z) (g \<xi>) < 1 \<longrightarrow> cmod (f z) \<le> cmod (g \<xi>) + 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
828 |
apply (simp add: eventually_at) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
829 |
apply (rule_tac x="\<delta>" in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
830 |
using \<delta> \<open>0 < \<delta>\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
831 |
apply (clarsimp simp:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
832 |
apply (drule_tac c=x in subsetD) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
833 |
apply (simp add: dist_commute) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
834 |
by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
835 |
have "continuous_on (interior S) g" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
836 |
by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
837 |
then have "\<And>x. x \<in> interior S \<Longrightarrow> (g \<longlongrightarrow> g x) (at x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
838 |
using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
839 |
then have "(g \<longlongrightarrow> g \<xi>) (at \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
840 |
by (simp add: \<xi>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
841 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
842 |
apply (rule_tac x="norm(g \<xi>) + 1" in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
843 |
apply (rule eventually_mp [OF * tendstoD [where e=1]], auto) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
844 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
845 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
846 |
moreover have "?Q" if "\<forall>\<^sub>F z in at \<xi>. cmod (f z) \<le> B" for B |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
847 |
by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
848 |
moreover have "?P" if "(\<lambda>z. (z - \<xi>) * f z) \<midarrow>\<xi>\<rightarrow> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
849 |
proof - |
63040 | 850 |
define h where [abs_def]: "h z = (z - \<xi>)^2 * f z" for z |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
851 |
have h0: "(h has_field_derivative 0) (at \<xi>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
852 |
apply (simp add: h_def Derivative.DERIV_within_iff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
853 |
apply (rule Lim_transform_within [OF that, of 1]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
854 |
apply (auto simp: divide_simps power2_eq_square) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
855 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
856 |
have holh: "h holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
857 |
proof (simp add: holomorphic_on_def, clarify) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
858 |
fix z assume "z \<in> S" |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
859 |
show "h field_differentiable at z within S" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
860 |
proof (cases "z = \<xi>") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
861 |
case True then show ?thesis |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
862 |
using field_differentiable_at_within field_differentiable_def h0 by blast |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
863 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
864 |
case False |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
865 |
then have "f field_differentiable at z within S" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
866 |
using holomorphic_onD [OF holf, of z] \<open>z \<in> S\<close> |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
867 |
unfolding field_differentiable_def DERIV_within_iff |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
868 |
by (force intro: exI [where x="dist \<xi> z"] elim: Lim_transform_within_set [unfolded eventually_at]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
869 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
870 |
by (simp add: h_def power2_eq_square derivative_intros) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
871 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
872 |
qed |
63040 | 873 |
define g where [abs_def]: "g z = (if z = \<xi> then deriv h \<xi> else (h z - h \<xi>) / (z - \<xi>))" for z |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
874 |
have holg: "g holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
875 |
unfolding g_def by (rule pole_lemma [OF holh \<xi>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
876 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
877 |
apply (rule_tac x="\<lambda>z. if z = \<xi> then deriv g \<xi> else (g z - g \<xi>)/(z - \<xi>)" in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
878 |
apply (rule conjI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
879 |
apply (rule pole_lemma [OF holg \<xi>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
880 |
apply (auto simp: g_def power2_eq_square divide_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
881 |
using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
882 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
883 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
884 |
ultimately show "?P = ?Q" and "?P = ?R" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
885 |
by meson+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
886 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
887 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
888 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
889 |
proposition pole_at_infinity: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
890 |
assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \<longlongrightarrow> l) at_infinity" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
891 |
obtains a n where "\<And>z. f z = (\<Sum>i\<le>n. a i * z^i)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
892 |
proof (cases "l = 0") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
893 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
894 |
with tendsto_inverse [OF lim] show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
895 |
apply (rule_tac a="(\<lambda>n. inverse l)" and n=0 in that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
896 |
apply (simp add: Liouville_weak [OF holf, of "inverse l"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
897 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
898 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
899 |
case True |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
900 |
then have [simp]: "l = 0" . |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
901 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
902 |
proof (cases "\<exists>r. 0 < r \<and> (\<forall>z \<in> ball 0 r - {0}. f(inverse z) \<noteq> 0)") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
903 |
case True |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
904 |
then obtain r where "0 < r" and r: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> f(inverse z) \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
905 |
by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
906 |
have 1: "inverse \<circ> f \<circ> inverse holomorphic_on ball 0 r - {0}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
907 |
by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
908 |
have 2: "0 \<in> interior (ball 0 r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
909 |
using \<open>0 < r\<close> by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
910 |
have "\<exists>B. 0<B \<and> eventually (\<lambda>z. cmod ((inverse \<circ> f \<circ> inverse) z) \<le> B) (at 0)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
911 |
apply (rule exI [where x=1]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
912 |
apply (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
913 |
using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
914 |
apply (rule eventually_mono) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
915 |
apply (simp add: dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
916 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
917 |
with holomorphic_on_extend_bounded [OF 1 2] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
918 |
obtain g where holg: "g holomorphic_on ball 0 r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
919 |
and geq: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> g z = (inverse \<circ> f \<circ> inverse) z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
920 |
by meson |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
921 |
have ifi0: "(inverse \<circ> f \<circ> inverse) \<midarrow>0\<rightarrow> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
922 |
using \<open>l = 0\<close> lim lim_at_infinity_0 by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
923 |
have g2g0: "g \<midarrow>0\<rightarrow> g 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
924 |
using \<open>0 < r\<close> centre_in_ball continuous_at continuous_on_eq_continuous_at holg |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
925 |
by (blast intro: holomorphic_on_imp_continuous_on) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
926 |
have g2g1: "g \<midarrow>0\<rightarrow> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
927 |
apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
928 |
using \<open>0 < r\<close> by (auto simp: geq) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
929 |
have [simp]: "g 0 = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
930 |
by (rule tendsto_unique [OF _ g2g0 g2g1]) simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
931 |
have "ball 0 r - {0::complex} \<noteq> {}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
932 |
using \<open>0 < r\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
933 |
apply (clarsimp simp: ball_def dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
934 |
apply (drule_tac c="of_real r/2" in subsetD, auto) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
935 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
936 |
then obtain w::complex where "w \<noteq> 0" and w: "norm w < r" by force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
937 |
then have "g w \<noteq> 0" by (simp add: geq r) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
938 |
obtain B n e where "0 < B" "0 < e" "e \<le> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
939 |
and leg: "\<And>w. norm w < e \<Longrightarrow> B * cmod w ^ n \<le> cmod (g w)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
940 |
apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
941 |
using \<open>0 < r\<close> w \<open>g w \<noteq> 0\<close> by (auto simp: ball_subset_ball_iff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
942 |
have "cmod (f z) \<le> cmod z ^ n / B" if "2/e \<le> cmod z" for z |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
943 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
944 |
have ize: "inverse z \<in> ball 0 e - {0}" using that \<open>0 < e\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
945 |
by (auto simp: norm_divide divide_simps algebra_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
946 |
then have [simp]: "z \<noteq> 0" and izr: "inverse z \<in> ball 0 r - {0}" using \<open>e \<le> r\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
947 |
by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
948 |
then have [simp]: "f z \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
949 |
using r [of "inverse z"] by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
950 |
have [simp]: "f z = inverse (g (inverse z))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
951 |
using izr geq [of "inverse z"] by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
952 |
show ?thesis using ize leg [of "inverse z"] \<open>0 < B\<close> \<open>0 < e\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
953 |
by (simp add: divide_simps norm_divide algebra_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
954 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
955 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
956 |
apply (rule_tac a = "\<lambda>k. (deriv ^^ k) f 0 / (fact k)" and n=n in that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
957 |
apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
958 |
apply (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
959 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
960 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
961 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
962 |
then have fi0: "\<And>r. r > 0 \<Longrightarrow> \<exists>z\<in>ball 0 r - {0}. f (inverse z) = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
963 |
by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
964 |
have fz0: "f z = 0" if "0 < r" and lt1: "\<And>x. x \<noteq> 0 \<Longrightarrow> cmod x < r \<Longrightarrow> inverse (cmod (f (inverse x))) < 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
965 |
for z r |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
966 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
967 |
have f0: "(f \<longlongrightarrow> 0) at_infinity" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
968 |
proof - |
62837 | 969 |
have DIM_complex[intro]: "2 \<le> DIM(complex)" \<comment>\<open>should not be necessary!\<close> |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
970 |
by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
971 |
have "continuous_on (inverse ` (ball 0 r - {0})) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
972 |
using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
973 |
then have "connected ((f \<circ> inverse) ` (ball 0 r - {0}))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
974 |
apply (intro connected_continuous_image continuous_intros) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
975 |
apply (force intro: connected_punctured_ball)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
976 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
977 |
then have "\<lbrakk>w \<noteq> 0; cmod w < r\<rbrakk> \<Longrightarrow> f (inverse w) = 0" for w |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
978 |
apply (rule disjE [OF connected_closedD [where A = "{0}" and B = "- ball 0 1"]], auto) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
979 |
apply (metis (mono_tags, hide_lams) not_less_iff_gr_or_eq one_less_inverse lt1 zero_less_norm_iff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
980 |
using False \<open>0 < r\<close> apply fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
981 |
by (metis (no_types, hide_lams) Compl_iff IntI comp_apply empty_iff image_eqI insert_Diff_single insert_iff mem_ball_0 not_less_iff_gr_or_eq one_less_inverse that(2) zero_less_norm_iff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
982 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
983 |
apply (simp add: lim_at_infinity_0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
984 |
apply (rule Lim_eventually) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
985 |
apply (simp add: eventually_at) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
986 |
apply (rule_tac x=r in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
987 |
apply (simp add: \<open>0 < r\<close> dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
988 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
989 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
990 |
obtain w where "w \<in> ball 0 r - {0}" and "f (inverse w) = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
991 |
using False \<open>0 < r\<close> by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
992 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
993 |
by (auto simp: f0 Liouville_weak [OF holf, of 0]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
994 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
995 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
996 |
apply (rule that [of "\<lambda>n. 0" 0]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
997 |
using lim [unfolded lim_at_infinity_0] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
998 |
apply (simp add: Lim_at dist_norm norm_inverse) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
999 |
apply (drule_tac x=1 in spec) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1000 |
using fz0 apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1001 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1002 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1003 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1004 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1005 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1006 |
subsection\<open>Entire proper functions are precisely the non-trivial polynomials\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1007 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1008 |
proposition proper_map_polyfun: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1009 |
fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1010 |
assumes "closed S" and "compact K" and c: "c i \<noteq> 0" "1 \<le> i" "i \<le> n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1011 |
shows "compact (S \<inter> {z. (\<Sum>i\<le>n. c i * z^i) \<in> K})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1012 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1013 |
obtain B where "B > 0" and B: "\<And>x. x \<in> K \<Longrightarrow> norm x \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1014 |
by (metis compact_imp_bounded \<open>compact K\<close> bounded_pos) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1015 |
have *: "norm x \<le> b" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1016 |
if "\<And>x. b \<le> norm x \<Longrightarrow> B + 1 \<le> norm (\<Sum>i\<le>n. c i * x ^ i)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1017 |
"(\<Sum>i\<le>n. c i * x ^ i) \<in> K" for b x |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1018 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1019 |
have "norm (\<Sum>i\<le>n. c i * x ^ i) \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1020 |
using B that by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1021 |
moreover have "\<not> B + 1 \<le> B" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1022 |
by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1023 |
ultimately show "norm x \<le> b" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1024 |
using that by (metis (no_types) less_eq_real_def not_less order_trans) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1025 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1026 |
have "bounded {z. (\<Sum>i\<le>n. c i * z ^ i) \<in> K}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1027 |
using polyfun_extremal [where c=c and B="B+1", OF c] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1028 |
by (auto simp: bounded_pos eventually_at_infinity_pos *) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1029 |
moreover have "closed {z. (\<Sum>i\<le>n. c i * z ^ i) \<in> K}" |
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63918
diff
changeset
|
1030 |
apply (intro allI continuous_closed_preimage_univ continuous_intros) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1031 |
using \<open>compact K\<close> compact_eq_bounded_closed by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1032 |
ultimately show ?thesis |
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62837
diff
changeset
|
1033 |
using closed_Int_compact [OF \<open>closed S\<close>] compact_eq_bounded_closed by blast |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1034 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1035 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1036 |
corollary proper_map_polyfun_univ: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1037 |
fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1038 |
assumes "compact K" "c i \<noteq> 0" "1 \<le> i" "i \<le> n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1039 |
shows "compact ({z. (\<Sum>i\<le>n. c i * z^i) \<in> K})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1040 |
using proper_map_polyfun [of UNIV K c i n] assms by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1041 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1042 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1043 |
proposition proper_map_polyfun_eq: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1044 |
assumes "f holomorphic_on UNIV" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1045 |
shows "(\<forall>k. compact k \<longrightarrow> compact {z. f z \<in> k}) \<longleftrightarrow> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1046 |
(\<exists>c n. 0 < n \<and> (c n \<noteq> 0) \<and> f = (\<lambda>z. \<Sum>i\<le>n. c i * z^i))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1047 |
(is "?lhs = ?rhs") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1048 |
proof |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1049 |
assume compf [rule_format]: ?lhs |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1050 |
have 2: "\<exists>k. 0 < k \<and> a k \<noteq> 0 \<and> f = (\<lambda>z. \<Sum>i \<le> k. a i * z ^ i)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1051 |
if "\<And>z. f z = (\<Sum>i\<le>n. a i * z ^ i)" for a n |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1052 |
proof (cases "\<forall>i\<le>n. 0<i \<longrightarrow> a i = 0") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1053 |
case True |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1054 |
then have [simp]: "\<And>z. f z = a 0" |
64267 | 1055 |
by (simp add: that sum_atMost_shift) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1056 |
have False using compf [of "{a 0}"] by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1057 |
then show ?thesis .. |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1058 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1059 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1060 |
then obtain k where k: "0 < k" "k\<le>n" "a k \<noteq> 0" by force |
63040 | 1061 |
define m where "m = (GREATEST k. k\<le>n \<and> a k \<noteq> 0)" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1062 |
have m: "m\<le>n \<and> a m \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1063 |
unfolding m_def |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1064 |
apply (rule GreatestI [where b = "Suc n"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1065 |
using k apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1066 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1067 |
have [simp]: "a i = 0" if "m < i" "i \<le> n" for i |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1068 |
using Greatest_le [where b = "Suc n" and P = "\<lambda>k. k\<le>n \<and> a k \<noteq> 0"] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1069 |
using m_def not_le that by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1070 |
have "k \<le> m" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1071 |
unfolding m_def |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1072 |
apply (rule Greatest_le [where b = "Suc n"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1073 |
using k apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1074 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1075 |
with k m show ?thesis |
64267 | 1076 |
by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1077 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1078 |
have "((inverse \<circ> f) \<longlongrightarrow> 0) at_infinity" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1079 |
proof (rule Lim_at_infinityI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1080 |
fix e::real assume "0 < e" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1081 |
with compf [of "cball 0 (inverse e)"] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1082 |
show "\<exists>B. \<forall>x. B \<le> cmod x \<longrightarrow> dist ((inverse \<circ> f) x) 0 \<le> e" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1083 |
apply (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1084 |
apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1085 |
apply (rule_tac x="b+1" in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1086 |
apply (metis inverse_inverse_eq less_add_same_cancel2 less_imp_inverse_less add.commute not_le not_less_iff_gr_or_eq order_trans zero_less_one) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1087 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1088 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1089 |
then show ?rhs |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1090 |
apply (rule pole_at_infinity [OF assms]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1091 |
using 2 apply blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1092 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1093 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1094 |
assume ?rhs |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1095 |
then obtain c n where "0 < n" "c n \<noteq> 0" "f = (\<lambda>z. \<Sum>i\<le>n. c i * z ^ i)" by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1096 |
then have "compact {z. f z \<in> k}" if "compact k" for k |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1097 |
by (auto intro: proper_map_polyfun_univ [OF that]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1098 |
then show ?lhs by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1099 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1100 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1101 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1102 |
subsection\<open>Relating invertibility and nonvanishing of derivative\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1103 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1104 |
proposition has_complex_derivative_locally_injective: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1105 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1106 |
and S: "\<xi> \<in> S" "open S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1107 |
and dnz: "deriv f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1108 |
obtains r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1109 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1110 |
have *: "\<exists>d>0. \<forall>x. dist \<xi> x < d \<longrightarrow> onorm (\<lambda>v. deriv f x * v - deriv f \<xi> * v) < e" if "e > 0" for e |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1111 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1112 |
have contdf: "continuous_on S (deriv f)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1113 |
by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \<open>open S\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1114 |
obtain \<delta> where "\<delta>>0" and \<delta>: "\<And>x. \<lbrakk>x \<in> S; dist x \<xi> \<le> \<delta>\<rbrakk> \<Longrightarrow> cmod (deriv f x - deriv f \<xi>) \<le> e/2" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1115 |
using continuous_onE [OF contdf \<open>\<xi> \<in> S\<close>, of "e/2"] \<open>0 < e\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1116 |
by (metis dist_complex_def half_gt_zero less_imp_le) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1117 |
obtain \<epsilon> where "\<epsilon>>0" "ball \<xi> \<epsilon> \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1118 |
by (metis openE [OF \<open>open S\<close> \<open>\<xi> \<in> S\<close>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1119 |
with \<open>\<delta>>0\<close> have "\<exists>\<delta>>0. \<forall>x. dist \<xi> x < \<delta> \<longrightarrow> onorm (\<lambda>v. deriv f x * v - deriv f \<xi> * v) \<le> e/2" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1120 |
apply (rule_tac x="min \<delta> \<epsilon>" in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1121 |
apply (intro conjI allI impI Operator_Norm.onorm_le) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1122 |
apply (simp add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1123 |
apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1124 |
apply (rule mult_right_mono [OF \<delta>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1125 |
apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \<delta>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1126 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1127 |
with \<open>e>0\<close> show ?thesis by force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1128 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1129 |
have "inj (op * (deriv f \<xi>))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1130 |
using dnz by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1131 |
then obtain g' where g': "linear g'" "g' \<circ> op * (deriv f \<xi>) = id" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1132 |
using linear_injective_left_inverse [of "op * (deriv f \<xi>)"] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1133 |
by (auto simp: linear_times) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1134 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1135 |
apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\<lambda>z h. deriv f z * h" and g' = g']) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1136 |
using g' * |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1137 |
apply (simp_all add: linear_conv_bounded_linear that) |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1138 |
using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1139 |
holomorphic_on_imp_differentiable_at \<open>open S\<close> apply blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1140 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1141 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1142 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1143 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1144 |
proposition has_complex_derivative_locally_invertible: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1145 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1146 |
and S: "\<xi> \<in> S" "open S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1147 |
and dnz: "deriv f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1148 |
obtains r where "r > 0" "ball \<xi> r \<subseteq> S" "open (f ` (ball \<xi> r))" "inj_on f (ball \<xi> r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1149 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1150 |
obtain r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1151 |
by (blast intro: that has_complex_derivative_locally_injective [OF assms]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1152 |
then have \<xi>: "\<xi> \<in> ball \<xi> r" by simp |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1153 |
then have nc: "~ f constant_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1154 |
using \<open>inj_on f (ball \<xi> r)\<close> injective_not_constant by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1155 |
have holf': "f holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1156 |
using \<open>ball \<xi> r \<subseteq> S\<close> holf holomorphic_on_subset by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1157 |
have "open (f ` ball \<xi> r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1158 |
apply (rule open_mapping_thm [OF holf']) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1159 |
using nc apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1160 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1161 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1162 |
using \<open>0 < r\<close> \<open>ball \<xi> r \<subseteq> S\<close> \<open>inj_on f (ball \<xi> r)\<close> that by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1163 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1164 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1165 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1166 |
proposition holomorphic_injective_imp_regular: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1167 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1168 |
and "open S" and injf: "inj_on f S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1169 |
and "\<xi> \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1170 |
shows "deriv f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1171 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1172 |
obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1173 |
have holf': "f holomorphic_on ball \<xi> r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1174 |
using \<open>ball \<xi> r \<subseteq> S\<close> holf holomorphic_on_subset by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1175 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1176 |
proof (cases "\<forall>n>0. (deriv ^^ n) f \<xi> = 0") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1177 |
case True |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1178 |
have fcon: "f w = f \<xi>" if "w \<in> ball \<xi> r" for w |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1179 |
apply (rule holomorphic_fun_eq_const_on_connected [OF holf']) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1180 |
using True \<open>0 < r\<close> that by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1181 |
have False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1182 |
using fcon [of "\<xi> + r/2"] \<open>0 < r\<close> r injf unfolding inj_on_def |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1183 |
by (metis \<open>\<xi> \<in> S\<close> contra_subsetD dist_commute fcon mem_ball perfect_choose_dist) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1184 |
then show ?thesis .. |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1185 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1186 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1187 |
then obtain n0 where n0: "n0 > 0 \<and> (deriv ^^ n0) f \<xi> \<noteq> 0" by blast |
63040 | 1188 |
define n where [abs_def]: "n = (LEAST n. n > 0 \<and> (deriv ^^ n) f \<xi> \<noteq> 0)" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1189 |
have n_ne: "n > 0" "(deriv ^^ n) f \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1190 |
using def_LeastI [OF n_def n0] by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1191 |
have n_min: "\<And>k. 0 < k \<Longrightarrow> k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1192 |
using def_Least_le [OF n_def] not_le by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1193 |
obtain g \<delta> where "0 < \<delta>" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1194 |
and holg: "g holomorphic_on ball \<xi> \<delta>" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1195 |
and fd: "\<And>w. w \<in> ball \<xi> \<delta> \<Longrightarrow> f w - f \<xi> = ((w - \<xi>) * g w) ^ n" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1196 |
and gnz: "\<And>w. w \<in> ball \<xi> \<delta> \<Longrightarrow> g w \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1197 |
apply (rule holomorphic_factor_order_of_zero_strong [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> n_ne]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1198 |
apply (blast intro: n_min)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1199 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1200 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1201 |
proof (cases "n=1") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1202 |
case True |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1203 |
with n_ne show ?thesis by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1204 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1205 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1206 |
have holgw: "(\<lambda>w. (w - \<xi>) * g w) holomorphic_on ball \<xi> (min r \<delta>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1207 |
apply (rule holomorphic_intros)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1208 |
using holg by (simp add: holomorphic_on_subset subset_ball) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1209 |
have gd: "\<And>w. dist \<xi> w < \<delta> \<Longrightarrow> (g has_field_derivative deriv g w) (at w)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1210 |
using holg |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1211 |
by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1212 |
have *: "\<And>w. w \<in> ball \<xi> (min r \<delta>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1213 |
\<Longrightarrow> ((\<lambda>w. (w - \<xi>) * g w) has_field_derivative ((w - \<xi>) * deriv g w + g w)) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1214 |
(at w)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1215 |
by (rule gd derivative_eq_intros | simp)+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1216 |
have [simp]: "deriv (\<lambda>w. (w - \<xi>) * g w) \<xi> \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1217 |
using * [of \<xi>] \<open>0 < \<delta>\<close> \<open>0 < r\<close> by (simp add: DERIV_imp_deriv gnz) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1218 |
obtain T where "\<xi> \<in> T" "open T" and Tsb: "T \<subseteq> ball \<xi> (min r \<delta>)" and oimT: "open ((\<lambda>w. (w - \<xi>) * g w) ` T)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1219 |
apply (rule has_complex_derivative_locally_invertible [OF holgw, of \<xi>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1220 |
using \<open>0 < r\<close> \<open>0 < \<delta>\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1221 |
apply (simp_all add:) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1222 |
by (meson Topology_Euclidean_Space.open_ball centre_in_ball) |
63040 | 1223 |
define U where "U = (\<lambda>w. (w - \<xi>) * g w) ` T" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1224 |
have "open U" by (metis oimT U_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1225 |
have "0 \<in> U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1226 |
apply (auto simp: U_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1227 |
apply (rule image_eqI [where x = \<xi>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1228 |
apply (auto simp: \<open>\<xi> \<in> T\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1229 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1230 |
then obtain \<epsilon> where "\<epsilon>>0" and \<epsilon>: "cball 0 \<epsilon> \<subseteq> U" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1231 |
using \<open>open U\<close> open_contains_cball by blast |
63589 | 1232 |
then have "\<epsilon> * exp(2 * of_real pi * \<i> * (0/n)) \<in> cball 0 \<epsilon>" |
1233 |
"\<epsilon> * exp(2 * of_real pi * \<i> * (1/n)) \<in> cball 0 \<epsilon>" |
|
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1234 |
by (auto simp: norm_mult) |
63589 | 1235 |
with \<epsilon> have "\<epsilon> * exp(2 * of_real pi * \<i> * (0/n)) \<in> U" |
1236 |
"\<epsilon> * exp(2 * of_real pi * \<i> * (1/n)) \<in> U" by blast+ |
|
1237 |
then obtain y0 y1 where "y0 \<in> T" and y0: "(y0 - \<xi>) * g y0 = \<epsilon> * exp(2 * of_real pi * \<i> * (0/n))" |
|
1238 |
and "y1 \<in> T" and y1: "(y1 - \<xi>) * g y1 = \<epsilon> * exp(2 * of_real pi * \<i> * (1/n))" |
|
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1239 |
by (auto simp: U_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1240 |
then have "y0 \<in> ball \<xi> \<delta>" "y1 \<in> ball \<xi> \<delta>" using Tsb by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1241 |
moreover have "y0 \<noteq> y1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1242 |
using y0 y1 \<open>\<epsilon> > 0\<close> complex_root_unity_eq_1 [of n 1] \<open>n > 0\<close> False by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1243 |
moreover have "T \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1244 |
by (meson Tsb min.cobounded1 order_trans r subset_ball) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1245 |
ultimately have False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1246 |
using inj_onD [OF injf, of y0 y1] \<open>y0 \<in> T\<close> \<open>y1 \<in> T\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1247 |
using fd [of y0] fd [of y1] complex_root_unity [of n 1] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1248 |
apply (simp add: y0 y1 power_mult_distrib) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1249 |
apply (force simp: algebra_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1250 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1251 |
then show ?thesis .. |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1252 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1253 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1254 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1255 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1256 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1257 |
text\<open>Hence a nice clean inverse function theorem\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1258 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1259 |
proposition holomorphic_has_inverse: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1260 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1261 |
and "open S" and injf: "inj_on f S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1262 |
obtains g where "g holomorphic_on (f ` S)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1263 |
"\<And>z. z \<in> S \<Longrightarrow> deriv f z * deriv g (f z) = 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1264 |
"\<And>z. z \<in> S \<Longrightarrow> g(f z) = z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1265 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1266 |
have ofs: "open (f ` S)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1267 |
by (rule open_mapping_thm3 [OF assms]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1268 |
have contf: "continuous_on S f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1269 |
by (simp add: holf holomorphic_on_imp_continuous_on) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1270 |
have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \<in> S" for z |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1271 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1272 |
have 1: "(f has_field_derivative deriv f z) (at z)" |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1273 |
using DERIV_deriv_iff_field_differentiable \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_on_imp_differentiable_at |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1274 |
by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1275 |
have 2: "deriv f z \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1276 |
using \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_injective_imp_regular injf by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1277 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1278 |
apply (rule has_complex_derivative_inverse_strong [OF 1 2 \<open>open S\<close> \<open>z \<in> S\<close>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1279 |
apply (simp add: holf holomorphic_on_imp_continuous_on) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1280 |
by (simp add: injf the_inv_into_f_f) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1281 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1282 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1283 |
proof |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1284 |
show "the_inv_into S f holomorphic_on f ` S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1285 |
by (simp add: holomorphic_on_open ofs) (blast intro: *) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1286 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1287 |
fix z assume "z \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1288 |
have "deriv f z \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1289 |
using \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_injective_imp_regular injf by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1290 |
then show "deriv f z * deriv (the_inv_into S f) (f z) = 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1291 |
using * [OF \<open>z \<in> S\<close>] by (simp add: DERIV_imp_deriv) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1292 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1293 |
fix z assume "z \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1294 |
show "the_inv_into S f (f z) = z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1295 |
by (simp add: \<open>z \<in> S\<close> injf the_inv_into_f_f) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1296 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1297 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1298 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1299 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1300 |
subsection\<open>The Schwarz Lemma\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1301 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1302 |
lemma Schwarz1: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1303 |
assumes holf: "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1304 |
and contf: "continuous_on (closure S) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1305 |
and S: "open S" "connected S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1306 |
and boS: "bounded S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1307 |
and "S \<noteq> {}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1308 |
obtains w where "w \<in> frontier S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1309 |
"\<And>z. z \<in> closure S \<Longrightarrow> norm (f z) \<le> norm (f w)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1310 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1311 |
have connf: "continuous_on (closure S) (norm o f)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1312 |
using contf continuous_on_compose continuous_on_norm_id by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1313 |
have coc: "compact (closure S)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1314 |
by (simp add: \<open>bounded S\<close> bounded_closure compact_eq_bounded_closed) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1315 |
then obtain x where x: "x \<in> closure S" and xmax: "\<And>z. z \<in> closure S \<Longrightarrow> norm(f z) \<le> norm(f x)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1316 |
apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1317 |
using \<open>S \<noteq> {}\<close> apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1318 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1319 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1320 |
proof (cases "x \<in> frontier S") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1321 |
case True |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1322 |
then show ?thesis using that xmax by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1323 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1324 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1325 |
then have "x \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1326 |
using \<open>open S\<close> frontier_def interior_eq x by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1327 |
then have "f constant_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1328 |
apply (rule maximum_modulus_principle [OF holf S \<open>open S\<close> order_refl]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1329 |
using closure_subset apply (blast intro: xmax) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1330 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1331 |
then have "f constant_on (closure S)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1332 |
by (rule constant_on_closureI [OF _ contf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1333 |
then obtain c where c: "\<And>x. x \<in> closure S \<Longrightarrow> f x = c" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1334 |
by (meson constant_on_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1335 |
obtain w where "w \<in> frontier S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1336 |
by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1337 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1338 |
by (simp add: c frontier_def that) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1339 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1340 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1341 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1342 |
lemma Schwarz2: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1343 |
"\<lbrakk>f holomorphic_on ball 0 r; |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1344 |
0 < s; ball w s \<subseteq> ball 0 r; |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1345 |
\<And>z. norm (w-z) < s \<Longrightarrow> norm(f z) \<le> norm(f w)\<rbrakk> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1346 |
\<Longrightarrow> f constant_on ball 0 r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1347 |
by (rule maximum_modulus_principle [where U = "ball w s" and \<xi> = w]) (simp_all add: dist_norm) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1348 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1349 |
lemma Schwarz3: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1350 |
assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1351 |
obtains h where "h holomorphic_on (ball 0 r)" and "\<And>z. norm z < r \<Longrightarrow> f z = z * (h z)" and "deriv f 0 = h 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1352 |
proof - |
63040 | 1353 |
define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1354 |
have d0: "deriv f 0 = h 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1355 |
by (simp add: h_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1356 |
moreover have "h holomorphic_on (ball 0 r)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1357 |
by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1358 |
moreover have "norm z < r \<Longrightarrow> f z = z * h z" for z |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1359 |
by (simp add: h_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1360 |
ultimately show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1361 |
using that by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1362 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1363 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1364 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1365 |
proposition Schwarz_Lemma: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1366 |
assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1367 |
and no: "\<And>z. norm z < 1 \<Longrightarrow> norm (f z) < 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1368 |
and \<xi>: "norm \<xi> < 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1369 |
shows "norm (f \<xi>) \<le> norm \<xi>" and "norm(deriv f 0) \<le> 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1370 |
and "((\<exists>z. norm z < 1 \<and> z \<noteq> 0 \<and> norm(f z) = norm z) \<or> norm(deriv f 0) = 1) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1371 |
\<Longrightarrow> \<exists>\<alpha>. (\<forall>z. norm z < 1 \<longrightarrow> f z = \<alpha> * z) \<and> norm \<alpha> = 1" (is "?P \<Longrightarrow> ?Q") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1372 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1373 |
obtain h where holh: "h holomorphic_on (ball 0 1)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1374 |
and fz_eq: "\<And>z. norm z < 1 \<Longrightarrow> f z = z * (h z)" and df0: "deriv f 0 = h 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1375 |
by (rule Schwarz3 [OF holf]) auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1376 |
have noh_le: "norm (h z) \<le> 1" if z: "norm z < 1" for z |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1377 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1378 |
have "norm (h z) < a" if a: "1 < a" for a |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1379 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1380 |
have "max (inverse a) (norm z) < 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1381 |
using z a by (simp_all add: inverse_less_1_iff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1382 |
then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1383 |
using Rats_dense_in_real by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1384 |
then have nzr: "norm z < r" and ira: "inverse r < a" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1385 |
using z a less_imp_inverse_less by force+ |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1386 |
then have "0 < r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1387 |
by (meson norm_not_less_zero not_le order.strict_trans2) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1388 |
have holh': "h holomorphic_on ball 0 r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1389 |
by (meson holh \<open>r < 1\<close> holomorphic_on_subset less_eq_real_def subset_ball) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1390 |
have conth': "continuous_on (cball 0 r) h" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1391 |
by (meson \<open>r < 1\<close> dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1392 |
obtain w where w: "norm w = r" and lenw: "\<And>z. norm z < r \<Longrightarrow> norm(h z) \<le> norm(h w)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1393 |
apply (rule Schwarz1 [OF holh']) using conth' \<open>0 < r\<close> by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1394 |
have "h w = f w / w" using fz_eq \<open>r < 1\<close> nzr w by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1395 |
then have "cmod (h z) < inverse r" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1396 |
by (metis \<open>0 < r\<close> \<open>r < 1\<close> divide_strict_right_mono inverse_eq_divide |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1397 |
le_less_trans lenw no norm_divide nzr w) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1398 |
then show ?thesis using ira by linarith |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1399 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1400 |
then show "norm (h z) \<le> 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1401 |
using not_le by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1402 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1403 |
show "cmod (f \<xi>) \<le> cmod \<xi>" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1404 |
proof (cases "\<xi> = 0") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1405 |
case True then show ?thesis by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1406 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1407 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1408 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1409 |
by (simp add: noh_le fz_eq \<xi> mult_left_le norm_mult) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1410 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1411 |
show no_df0: "norm(deriv f 0) \<le> 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1412 |
by (simp add: \<open>\<And>z. cmod z < 1 \<Longrightarrow> cmod (h z) \<le> 1\<close> df0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1413 |
show "?Q" if "?P" |
63540 | 1414 |
using that |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1415 |
proof |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1416 |
assume "\<exists>z. cmod z < 1 \<and> z \<noteq> 0 \<and> cmod (f z) = cmod z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1417 |
then obtain \<gamma> where \<gamma>: "cmod \<gamma> < 1" "\<gamma> \<noteq> 0" "cmod (f \<gamma>) = cmod \<gamma>" by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1418 |
then have [simp]: "norm (h \<gamma>) = 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1419 |
by (simp add: fz_eq norm_mult) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1420 |
have "ball \<gamma> (1 - cmod \<gamma>) \<subseteq> ball 0 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1421 |
by (simp add: ball_subset_ball_iff) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1422 |
moreover have "\<And>z. cmod (\<gamma> - z) < 1 - cmod \<gamma> \<Longrightarrow> cmod (h z) \<le> cmod (h \<gamma>)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1423 |
apply (simp add: algebra_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1424 |
by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1425 |
ultimately obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1426 |
using Schwarz2 [OF holh, of "1 - norm \<gamma>" \<gamma>, unfolded constant_on_def] \<gamma> by auto |
63540 | 1427 |
then have "norm c = 1" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1428 |
using \<gamma> by force |
63540 | 1429 |
with c show ?thesis |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1430 |
using fz_eq by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1431 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1432 |
assume [simp]: "cmod (deriv f 0) = 1" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1433 |
then obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1434 |
using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1435 |
by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1436 |
moreover have "norm c = 1" using df0 c by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1437 |
ultimately show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1438 |
using fz_eq by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1439 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1440 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1441 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1442 |
subsection\<open>The Schwarz reflection principle\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1443 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1444 |
lemma hol_pal_lem0: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1445 |
assumes "d \<bullet> a \<le> k" "k \<le> d \<bullet> b" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1446 |
obtains c where |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1447 |
"c \<in> closed_segment a b" "d \<bullet> c = k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1448 |
"\<And>z. z \<in> closed_segment a c \<Longrightarrow> d \<bullet> z \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1449 |
"\<And>z. z \<in> closed_segment c b \<Longrightarrow> k \<le> d \<bullet> z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1450 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1451 |
obtain c where cin: "c \<in> closed_segment a b" and keq: "k = d \<bullet> c" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1452 |
using connected_ivt_hyperplane [of "closed_segment a b" a b d k] |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1453 |
by (auto simp: assms) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1454 |
have "closed_segment a c \<subseteq> {z. d \<bullet> z \<le> k}" "closed_segment c b \<subseteq> {z. k \<le> d \<bullet> z}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1455 |
unfolding segment_convex_hull using assms keq |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1456 |
by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1457 |
then show ?thesis using cin that by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1458 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1459 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1460 |
lemma hol_pal_lem1: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1461 |
assumes "convex S" "open S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1462 |
and abc: "a \<in> S" "b \<in> S" "c \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1463 |
"d \<noteq> 0" and lek: "d \<bullet> a \<le> k" "d \<bullet> b \<le> k" "d \<bullet> c \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1464 |
and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1465 |
and contf: "continuous_on S f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1466 |
shows "contour_integral (linepath a b) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1467 |
contour_integral (linepath b c) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1468 |
contour_integral (linepath c a) f = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1469 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1470 |
have "interior (convex hull {a, b, c}) \<subseteq> interior(S \<inter> {x. d \<bullet> x \<le> k})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1471 |
apply (rule interior_mono) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1472 |
apply (rule hull_minimal) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1473 |
apply (simp add: abc lek) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1474 |
apply (rule convex_Int [OF \<open>convex S\<close> convex_halfspace_le]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1475 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1476 |
also have "... \<subseteq> {z \<in> S. d \<bullet> z < k}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1477 |
by (force simp: interior_open [OF \<open>open S\<close>] \<open>d \<noteq> 0\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1478 |
finally have *: "interior (convex hull {a, b, c}) \<subseteq> {z \<in> S. d \<bullet> z < k}" . |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1479 |
have "continuous_on (convex hull {a,b,c}) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1480 |
using \<open>convex S\<close> contf abc continuous_on_subset subset_hull |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1481 |
by fastforce |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1482 |
moreover have "f holomorphic_on interior (convex hull {a,b,c})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1483 |
by (rule holomorphic_on_subset [OF holf1 *]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1484 |
ultimately show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1485 |
using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3 |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1486 |
by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1487 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1488 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1489 |
lemma hol_pal_lem2: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1490 |
assumes S: "convex S" "open S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1491 |
and abc: "a \<in> S" "b \<in> S" "c \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1492 |
and "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" "d \<bullet> b \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1493 |
and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1494 |
and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1495 |
and contf: "continuous_on S f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1496 |
shows "contour_integral (linepath a b) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1497 |
contour_integral (linepath b c) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1498 |
contour_integral (linepath c a) f = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1499 |
proof (cases "d \<bullet> c \<le> k") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1500 |
case True show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1501 |
by (rule hol_pal_lem1 [OF S abc \<open>d \<noteq> 0\<close> lek True holf1 contf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1502 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1503 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1504 |
then have "d \<bullet> c > k" by force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1505 |
obtain a' where a': "a' \<in> closed_segment b c" and "d \<bullet> a' = k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1506 |
and ba': "\<And>z. z \<in> closed_segment b a' \<Longrightarrow> d \<bullet> z \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1507 |
and a'c: "\<And>z. z \<in> closed_segment a' c \<Longrightarrow> k \<le> d \<bullet> z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1508 |
apply (rule hol_pal_lem0 [of d b k c, OF \<open>d \<bullet> b \<le> k\<close>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1509 |
using False by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1510 |
obtain b' where b': "b' \<in> closed_segment a c" and "d \<bullet> b' = k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1511 |
and ab': "\<And>z. z \<in> closed_segment a b' \<Longrightarrow> d \<bullet> z \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1512 |
and b'c: "\<And>z. z \<in> closed_segment b' c \<Longrightarrow> k \<le> d \<bullet> z" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1513 |
apply (rule hol_pal_lem0 [of d a k c, OF \<open>d \<bullet> a \<le> k\<close>]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1514 |
using False by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1515 |
have a'b': "a' \<in> S \<and> b' \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1516 |
using a' abc b' convex_contains_segment \<open>convex S\<close> by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1517 |
have "continuous_on (closed_segment c a) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1518 |
by (meson abc contf continuous_on_subset convex_contains_segment \<open>convex S\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1519 |
then have 1: "contour_integral (linepath c a) f = |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1520 |
contour_integral (linepath c b') f + contour_integral (linepath b' a) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1521 |
apply (rule contour_integral_split_linepath) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1522 |
using b' by (simp add: closed_segment_commute) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1523 |
have "continuous_on (closed_segment b c) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1524 |
by (meson abc contf continuous_on_subset convex_contains_segment \<open>convex S\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1525 |
then have 2: "contour_integral (linepath b c) f = |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1526 |
contour_integral (linepath b a') f + contour_integral (linepath a' c) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1527 |
by (rule contour_integral_split_linepath [OF _ a']) |
62463
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
paulson <lp15@cam.ac.uk>
parents:
62408
diff
changeset
|
1528 |
have 3: "contour_integral (reversepath (linepath b' a')) f = |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1529 |
- contour_integral (linepath b' a') f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1530 |
by (rule contour_integral_reversepath [OF valid_path_linepath]) |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1531 |
have fcd_le: "f field_differentiable at x" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1532 |
if "x \<in> interior S \<and> x \<in> interior {x. d \<bullet> x \<le> k}" for x |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1533 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1534 |
have "f holomorphic_on S \<inter> {c. d \<bullet> c < k}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1535 |
by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1536 |
then have "\<exists>C D. x \<in> interior C \<inter> interior D \<and> f holomorphic_on interior C \<inter> interior D" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1537 |
using that |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1538 |
by (metis Collect_mem_eq Int_Collect \<open>d \<noteq> 0\<close> interior_halfspace_le interior_open \<open>open S\<close>) |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1539 |
then show "f field_differentiable at x" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1540 |
by (metis at_within_interior holomorphic_on_def interior_Int interior_interior) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1541 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1542 |
have ab_le: "\<And>x. x \<in> closed_segment a b \<Longrightarrow> d \<bullet> x \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1543 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1544 |
fix x :: complex |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1545 |
assume "x \<in> closed_segment a b" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1546 |
then have "\<And>C. x \<in> C \<or> b \<notin> C \<or> a \<notin> C \<or> \<not> convex C" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1547 |
by (meson contra_subsetD convex_contains_segment) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1548 |
then show "d \<bullet> x \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1549 |
by (metis lek convex_halfspace_le mem_Collect_eq) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1550 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1551 |
have "continuous_on (S \<inter> {x. d \<bullet> x \<le> k}) f" using contf |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1552 |
by (simp add: continuous_on_subset) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1553 |
then have "(f has_contour_integral 0) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1554 |
(linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1555 |
apply (rule Cauchy_theorem_convex [where k = "{}"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1556 |
apply (simp_all add: path_image_join convex_Int convex_halfspace_le \<open>convex S\<close> fcd_le ab_le |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1557 |
closed_segment_subset abc a'b' ba') |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1558 |
by (metis \<open>d \<bullet> a' = k\<close> \<open>d \<bullet> b' = k\<close> convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1559 |
then have 4: "contour_integral (linepath a b) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1560 |
contour_integral (linepath b a') f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1561 |
contour_integral (linepath a' b') f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1562 |
contour_integral (linepath b' a) f = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1563 |
by (rule has_chain_integral_chain_integral4) |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1564 |
have fcd_ge: "f field_differentiable at x" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1565 |
if "x \<in> interior S \<and> x \<in> interior {x. k \<le> d \<bullet> x}" for x |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1566 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1567 |
have f2: "f holomorphic_on S \<inter> {c. k < d \<bullet> c}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1568 |
by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1569 |
have f3: "interior S = S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1570 |
by (simp add: interior_open \<open>open S\<close>) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1571 |
then have "x \<in> S \<inter> interior {c. k \<le> d \<bullet> c}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1572 |
using that by simp |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1573 |
then show "f field_differentiable at x" |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1574 |
using f3 f2 unfolding holomorphic_on_def |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1575 |
by (metis (no_types) \<open>d \<noteq> 0\<close> at_within_interior interior_Int interior_halfspace_ge interior_interior) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1576 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1577 |
have "continuous_on (S \<inter> {x. k \<le> d \<bullet> x}) f" using contf |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1578 |
by (simp add: continuous_on_subset) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1579 |
then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1580 |
apply (rule Cauchy_theorem_convex [where k = "{}"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1581 |
apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \<open>convex S\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1582 |
fcd_ge closed_segment_subset abc a'b' a'c) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1583 |
by (metis \<open>d \<bullet> a' = k\<close> b'c closed_segment_commute convex_contains_segment |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1584 |
convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1585 |
then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1586 |
by (rule has_chain_integral_chain_integral3) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1587 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1588 |
using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1589 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1590 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1591 |
lemma hol_pal_lem3: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1592 |
assumes S: "convex S" "open S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1593 |
and abc: "a \<in> S" "b \<in> S" "c \<in> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1594 |
and "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1595 |
and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1596 |
and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1597 |
and contf: "continuous_on S f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1598 |
shows "contour_integral (linepath a b) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1599 |
contour_integral (linepath b c) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1600 |
contour_integral (linepath c a) f = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1601 |
proof (cases "d \<bullet> b \<le> k") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1602 |
case True show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1603 |
by (rule hol_pal_lem2 [OF S abc \<open>d \<noteq> 0\<close> lek True holf1 holf2 contf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1604 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1605 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1606 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1607 |
proof (cases "d \<bullet> c \<le> k") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1608 |
case True |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1609 |
have "contour_integral (linepath c a) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1610 |
contour_integral (linepath a b) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1611 |
contour_integral (linepath b c) f = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1612 |
by (rule hol_pal_lem2 [OF S \<open>c \<in> S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close> \<open>d \<noteq> 0\<close> \<open>d \<bullet> c \<le> k\<close> lek holf1 holf2 contf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1613 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1614 |
by (simp add: algebra_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1615 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1616 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1617 |
have "contour_integral (linepath b c) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1618 |
contour_integral (linepath c a) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1619 |
contour_integral (linepath a b) f = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1620 |
apply (rule hol_pal_lem2 [OF S \<open>b \<in> S\<close> \<open>c \<in> S\<close> \<open>a \<in> S\<close>, of "-d" "-k"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1621 |
using \<open>d \<noteq> 0\<close> \<open>\<not> d \<bullet> b \<le> k\<close> False by (simp_all add: holf1 holf2 contf) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1622 |
then show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1623 |
by (simp add: algebra_simps) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1624 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1625 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1626 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1627 |
lemma hol_pal_lem4: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1628 |
assumes S: "convex S" "open S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1629 |
and abc: "a \<in> S" "b \<in> S" "c \<in> S" and "d \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1630 |
and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1631 |
and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1632 |
and contf: "continuous_on S f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1633 |
shows "contour_integral (linepath a b) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1634 |
contour_integral (linepath b c) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1635 |
contour_integral (linepath c a) f = 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1636 |
proof (cases "d \<bullet> a \<le> k") |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1637 |
case True show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1638 |
by (rule hol_pal_lem3 [OF S abc \<open>d \<noteq> 0\<close> True holf1 holf2 contf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1639 |
next |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1640 |
case False |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1641 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1642 |
apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1643 |
using \<open>d \<noteq> 0\<close> False by (simp_all add: holf1 holf2 contf) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1644 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1645 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1646 |
proposition holomorphic_on_paste_across_line: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1647 |
assumes S: "open S" and "d \<noteq> 0" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1648 |
and holf1: "f holomorphic_on (S \<inter> {z. d \<bullet> z < k})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1649 |
and holf2: "f holomorphic_on (S \<inter> {z. k < d \<bullet> z})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1650 |
and contf: "continuous_on S f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1651 |
shows "f holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1652 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1653 |
have *: "\<exists>t. open t \<and> p \<in> t \<and> continuous_on t f \<and> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1654 |
(\<forall>a b c. convex hull {a, b, c} \<subseteq> t \<longrightarrow> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1655 |
contour_integral (linepath a b) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1656 |
contour_integral (linepath b c) f + |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1657 |
contour_integral (linepath c a) f = 0)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1658 |
if "p \<in> S" for p |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1659 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1660 |
obtain e where "e>0" and e: "ball p e \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1661 |
using \<open>p \<in> S\<close> openE S by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1662 |
then have "continuous_on (ball p e) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1663 |
using contf continuous_on_subset by blast |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1664 |
moreover have "f holomorphic_on {z. dist p z < e \<and> d \<bullet> z < k}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1665 |
apply (rule holomorphic_on_subset [OF holf1]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1666 |
using e by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1667 |
moreover have "f holomorphic_on {z. dist p z < e \<and> k < d \<bullet> z}" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1668 |
apply (rule holomorphic_on_subset [OF holf2]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1669 |
using e by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1670 |
ultimately show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1671 |
apply (rule_tac x="ball p e" in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1672 |
using \<open>e > 0\<close> e \<open>d \<noteq> 0\<close> |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1673 |
apply (simp add:, clarify) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1674 |
apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1675 |
apply (auto simp: subset_hull) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1676 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1677 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1678 |
show ?thesis |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1679 |
by (blast intro: * Morera_local_triangle analytic_imp_holomorphic) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1680 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1681 |
|
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1682 |
proposition Schwarz_reflection: |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1683 |
assumes "open S" and cnjs: "cnj ` S \<subseteq> S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1684 |
and holf: "f holomorphic_on (S \<inter> {z. 0 < Im z})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1685 |
and contf: "continuous_on (S \<inter> {z. 0 \<le> Im z}) f" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1686 |
and f: "\<And>z. \<lbrakk>z \<in> S; z \<in> \<real>\<rbrakk> \<Longrightarrow> (f z) \<in> \<real>" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1687 |
shows "(\<lambda>z. if 0 \<le> Im z then f z else cnj(f(cnj z))) holomorphic_on S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1688 |
proof - |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1689 |
have 1: "(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z))) holomorphic_on (S \<inter> {z. 0 < Im z})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1690 |
by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1691 |
have cont_cfc: "continuous_on (S \<inter> {z. Im z \<le> 0}) (cnj o f o cnj)" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1692 |
apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1693 |
using cnjs apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1694 |
done |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1695 |
have "cnj \<circ> f \<circ> cnj field_differentiable at x within S \<inter> {z. Im z < 0}" |
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1696 |
if "x \<in> S" "Im x < 0" "f field_differentiable at (cnj x) within S \<inter> {z. 0 < Im z}" for x |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1697 |
using that |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1698 |
apply (simp add: field_differentiable_def Derivative.DERIV_within_iff Lim_within dist_norm, clarify) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1699 |
apply (rule_tac x="cnj f'" in exI) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1700 |
apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1701 |
apply (drule_tac x="cnj xa" in bspec) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1702 |
using cnjs apply force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1703 |
apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1704 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1705 |
then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \<inter> {z. Im z < 0})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1706 |
using holf cnjs |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1707 |
by (force simp: holomorphic_on_def) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1708 |
have 2: "(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z))) holomorphic_on (S \<inter> {z. Im z < 0})" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1709 |
apply (rule iffD1 [OF holomorphic_cong [OF refl]]) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1710 |
using hol_cfc by auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1711 |
have [simp]: "(S \<inter> {z. 0 \<le> Im z}) \<union> (S \<inter> {z. Im z \<le> 0}) = S" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1712 |
by force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1713 |
have "continuous_on ((S \<inter> {z. 0 \<le> Im z}) \<union> (S \<inter> {z. Im z \<le> 0})) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1714 |
(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z)))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1715 |
apply (rule continuous_on_cases_local) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1716 |
using cont_cfc contf |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1717 |
apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge) |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1718 |
using f Reals_cnj_iff complex_is_Real_iff apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1719 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1720 |
then have 3: "continuous_on S (\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z)))" |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1721 |
by force |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1722 |
show ?thesis |
63589 | 1723 |
apply (rule holomorphic_on_paste_across_line [OF \<open>open S\<close>, of "- \<i>" _ 0]) |
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1724 |
using 1 2 3 |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1725 |
apply auto |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1726 |
done |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1727 |
qed |
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1728 |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1729 |
subsection\<open>Bloch's theorem\<close> |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1730 |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1731 |
lemma Bloch_lemma_0: |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1732 |
assumes holf: "f holomorphic_on cball 0 r" and "0 < r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1733 |
and [simp]: "f 0 = 0" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1734 |
and le: "\<And>z. norm z < r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f 0)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1735 |
shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \<subseteq> f ` ball 0 r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1736 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1737 |
have "sqrt 2 < 3/2" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1738 |
by (rule real_less_lsqrt) (auto simp: power2_eq_square) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1739 |
then have sq3: "0 < 3 - 2 * sqrt 2" by simp |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1740 |
show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1741 |
proof (cases "deriv f 0 = 0") |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1742 |
case True then show ?thesis by simp |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1743 |
next |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1744 |
case False |
63040 | 1745 |
define C where "C = 2 * norm(deriv f 0)" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1746 |
have "0 < C" using False by (simp add: C_def) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1747 |
have holf': "f holomorphic_on ball 0 r" using holf |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1748 |
using ball_subset_cball holomorphic_on_subset by blast |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1749 |
then have holdf': "deriv f holomorphic_on ball 0 r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1750 |
by (rule holomorphic_deriv [OF _ open_ball]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1751 |
have "Le1": "norm(deriv f z - deriv f 0) \<le> norm z / (r - norm z) * C" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1752 |
if "norm z < r" for z |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1753 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1754 |
have T1: "norm(deriv f z - deriv f 0) \<le> norm z / (R - norm z) * C" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1755 |
if R: "norm z < R" "R < r" for R |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1756 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1757 |
have "0 < R" using R |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1758 |
by (metis less_trans norm_zero zero_less_norm_iff) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1759 |
have df_le: "\<And>x. norm x < r \<Longrightarrow> norm (deriv f x) \<le> C" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1760 |
using le by (simp add: C_def) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1761 |
have hol_df: "deriv f holomorphic_on cball 0 R" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1762 |
apply (rule holomorphic_on_subset) using R holdf' by auto |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1763 |
have *: "((\<lambda>w. deriv f w / (w - z)) has_contour_integral 2 * pi * \<i> * deriv f z) (circlepath 0 R)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1764 |
if "norm z < R" for z |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1765 |
using \<open>0 < R\<close> that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"] |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1766 |
by (force simp: winding_number_circlepath) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1767 |
have **: "((\<lambda>x. deriv f x / (x - z) - deriv f x / x) has_contour_integral |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1768 |
of_real (2 * pi) * \<i> * (deriv f z - deriv f 0)) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1769 |
(circlepath 0 R)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1770 |
using has_contour_integral_diff [OF * [of z] * [of 0]] \<open>0 < R\<close> that |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1771 |
by (simp add: algebra_simps) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1772 |
have [simp]: "\<And>x. norm x = R \<Longrightarrow> x \<noteq> z" using that(1) by blast |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1773 |
have "norm (deriv f x / (x - z) - deriv f x / x) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1774 |
\<le> C * norm z / (R * (R - norm z))" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1775 |
if "norm x = R" for x |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1776 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1777 |
have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) = |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1778 |
norm (deriv f x) * norm z" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1779 |
by (simp add: norm_mult right_diff_distrib') |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1780 |
show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1781 |
using \<open>0 < R\<close> \<open>0 < C\<close> R that |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1782 |
apply (simp add: norm_mult norm_divide divide_simps) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1783 |
using df_le norm_triangle_ineq2 \<open>0 < C\<close> apply (auto intro!: mult_mono) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1784 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1785 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1786 |
then show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1787 |
using has_contour_integral_bound_circlepath |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1788 |
[OF **, of "C * norm z/(R*(R - norm z))"] |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1789 |
\<open>0 < R\<close> \<open>0 < C\<close> R |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1790 |
apply (simp add: norm_mult norm_divide) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1791 |
apply (simp add: divide_simps mult.commute) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1792 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1793 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1794 |
obtain r' where r': "norm z < r'" "r' < r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1795 |
using Rats_dense_in_real [of "norm z" r] \<open>norm z < r\<close> by blast |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1796 |
then have [simp]: "closure {r'<..<r} = {r'..r}" by simp |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1797 |
show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1798 |
apply (rule continuous_ge_on_closure |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1799 |
[where f = "\<lambda>r. norm z / (r - norm z) * C" and s = "{r'<..<r}", |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1800 |
OF _ _ T1]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1801 |
apply (intro continuous_intros) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1802 |
using that r' |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1803 |
apply (auto simp: not_le) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1804 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1805 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1806 |
have "*": "(norm z - norm z^2/(r - norm z)) * norm(deriv f 0) \<le> norm(f z)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1807 |
if r: "norm z < r" for z |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1808 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1809 |
have 1: "\<And>x. x \<in> ball 0 r \<Longrightarrow> |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1810 |
((\<lambda>z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1811 |
(at x within ball 0 r)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1812 |
by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+ |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1813 |
have 2: "closed_segment 0 z \<subseteq> ball 0 r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1814 |
by (metis \<open>0 < r\<close> convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1815 |
have 3: "(\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1816 |
apply (rule integrable_on_cmult_right [where 'b=real, simplified]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1817 |
apply (rule integrable_on_cdivide [where 'b=real, simplified]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1818 |
apply (rule integrable_on_cmult_left [where 'b=real, simplified]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1819 |
apply (rule ident_integrable_on) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1820 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1821 |
have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm z * norm z * x * C / (r - norm z)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1822 |
if x: "0 \<le> x" "x \<le> 1" for x |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1823 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1824 |
have [simp]: "x * norm z < r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1825 |
using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1826 |
have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \<le> norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1827 |
apply (rule Le1) using r x \<open>0 < r\<close> by simp |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1828 |
also have "... \<le> norm (x *\<^sub>R z) / (r - norm z) * C" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1829 |
using r x \<open>0 < r\<close> |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1830 |
apply (simp add: divide_simps) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1831 |
by (simp add: \<open>0 < C\<close> mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1832 |
finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm (x *\<^sub>R z) / (r - norm z) * C * norm z" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1833 |
by (rule mult_right_mono) simp |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1834 |
with x show ?thesis by (simp add: algebra_simps) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1835 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1836 |
have le_norm: "abc \<le> norm d - e \<Longrightarrow> norm(f - d) \<le> e \<Longrightarrow> abc \<le> norm f" for abc d e and f::complex |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1837 |
by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1838 |
have "norm (integral {0..1} (\<lambda>x. (deriv f (x *\<^sub>R z) - deriv f 0) * z)) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1839 |
\<le> integral {0..1} (\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1840 |
apply (rule integral_norm_bound_integral) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1841 |
using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1842 |
apply (simp add: has_contour_integral_linepath has_integral_integrable_integral) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1843 |
apply (rule 3) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1844 |
apply (simp add: norm_mult power2_eq_square 4) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1845 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1846 |
then have int_le: "norm (f z - deriv f 0 * z) \<le> (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1847 |
using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1848 |
apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1849 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1850 |
show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1851 |
apply (rule le_norm [OF _ int_le]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1852 |
using \<open>norm z < r\<close> |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1853 |
apply (simp add: power2_eq_square divide_simps C_def norm_mult) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1854 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1855 |
have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \<le> norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1856 |
by (simp add: linordered_field_class.sign_simps(38)) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1857 |
then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \<le> norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1858 |
by (simp add: linordered_field_class.sign_simps(38) mult.commute mult.left_commute) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1859 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1860 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1861 |
have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2) < 1" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1862 |
by (auto simp: sqrt2_less_2) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1863 |
have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1864 |
apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1865 |
apply (subst closure_ball) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1866 |
using \<open>0 < r\<close> mult_pos_pos sq201 |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1867 |
apply (auto simp: cball_subset_cball_iff) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1868 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1869 |
have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1870 |
apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1871 |
using \<open>0 < r\<close> mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1872 |
using False \<open>0 < r\<close> centre_in_ball holf' holomorphic_nonconstant by blast |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1873 |
have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) = |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1874 |
ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1875 |
by simp |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1876 |
also have "... \<subseteq> f ` ball 0 ((1 - sqrt 2 / 2) * r)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1877 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1878 |
have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \<le> norm (f z)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1879 |
if "norm z = (1 - sqrt 2 / 2) * r" for z |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1880 |
apply (rule order_trans [OF _ *]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1881 |
using \<open>0 < r\<close> |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1882 |
apply (simp_all add: field_simps power2_eq_square that) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1883 |
apply (simp add: mult.assoc [symmetric]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1884 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1885 |
show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1886 |
apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1887 |
using \<open>0 < r\<close> sq201 3 apply simp_all |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1888 |
using C_def \<open>0 < C\<close> sq3 apply force |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1889 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1890 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1891 |
also have "... \<subseteq> f ` ball 0 r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1892 |
apply (rule image_subsetI [OF imageI], simp) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1893 |
apply (erule less_le_trans) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1894 |
using \<open>0 < r\<close> apply (auto simp: field_simps) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1895 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1896 |
finally show ?thesis . |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1897 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1898 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1899 |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1900 |
lemma Bloch_lemma: |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1901 |
assumes holf: "f holomorphic_on cball a r" and "0 < r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1902 |
and le: "\<And>z. z \<in> ball a r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f a)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1903 |
shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \<subseteq> f ` ball a r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1904 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1905 |
have fz: "(\<lambda>z. f (a + z)) = f o (\<lambda>z. (a + z))" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1906 |
by (simp add: o_def) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1907 |
have hol0: "(\<lambda>z. f (a + z)) holomorphic_on cball 0 r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1908 |
unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+ |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1909 |
then have [simp]: "\<And>x. norm x < r \<Longrightarrow> (\<lambda>z. f (a + z)) field_differentiable at x" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1910 |
by (metis Topology_Euclidean_Space.open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel) |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1911 |
have [simp]: "\<And>z. norm z < r \<Longrightarrow> f field_differentiable at (a + z)" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1912 |
by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute) |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1913 |
then have [simp]: "f field_differentiable at a" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1914 |
by (metis add.comm_neutral \<open>0 < r\<close> norm_eq_zero) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1915 |
have hol1: "(\<lambda>z. f (a + z) - f a) holomorphic_on cball 0 r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1916 |
by (intro holomorphic_intros hol0) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1917 |
then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\<lambda>z. f (a + z) - f a) 0)) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1918 |
\<subseteq> (\<lambda>z. f (a + z) - f a) ` ball 0 r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1919 |
apply (rule Bloch_lemma_0) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1920 |
apply (simp_all add: \<open>0 < r\<close>) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1921 |
apply (simp add: fz complex_derivative_chain) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1922 |
apply (simp add: dist_norm le) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1923 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1924 |
then show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1925 |
apply clarify |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1926 |
apply (drule_tac c="x - f a" in subsetD) |
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
1927 |
apply (force simp: fz \<open>0 < r\<close> dist_norm complex_derivative_chain field_differentiable_compose)+ |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1928 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1929 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1930 |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1931 |
proposition Bloch_unit: |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1932 |
assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1933 |
obtains b r where "1/12 < r" "ball b r \<subseteq> f ` (ball a 1)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1934 |
proof - |
63040 | 1935 |
define r :: real where "r = 249/256" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1936 |
have "0 < r" "r < 1" by (auto simp: r_def) |
63040 | 1937 |
define g where "g z = deriv f z * of_real(r - norm(z - a))" for z |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1938 |
have "deriv f holomorphic_on ball a 1" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1939 |
by (rule holomorphic_deriv [OF holf open_ball]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1940 |
then have "continuous_on (ball a 1) (deriv f)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1941 |
using holomorphic_on_imp_continuous_on by blast |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1942 |
then have "continuous_on (cball a r) (deriv f)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1943 |
by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \<open>r < 1\<close>) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1944 |
then have "continuous_on (cball a r) g" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1945 |
by (simp add: g_def continuous_intros) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1946 |
then have 1: "compact (g ` cball a r)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1947 |
by (rule compact_continuous_image [OF _ compact_cball]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1948 |
have 2: "g ` cball a r \<noteq> {}" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1949 |
using \<open>r > 0\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1950 |
obtain p where pr: "p \<in> cball a r" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1951 |
and pge: "\<And>y. y \<in> cball a r \<Longrightarrow> norm (g y) \<le> norm (g p)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1952 |
using distance_attains_sup [OF 1 2, of 0] by force |
63040 | 1953 |
define t where "t = (r - norm(p - a)) / 2" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1954 |
have "norm (p - a) \<noteq> r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1955 |
using pge [of a] \<open>r > 0\<close> by (auto simp: g_def norm_mult) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1956 |
then have "norm (p - a) < r" using pr |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1957 |
by (simp add: norm_minus_commute dist_norm) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1958 |
then have "0 < t" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1959 |
by (simp add: t_def) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1960 |
have cpt: "cball p t \<subseteq> ball a r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1961 |
using \<open>0 < t\<close> by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1962 |
have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \<le> norm (deriv f p)" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1963 |
if "y \<in> cball a r" for y |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1964 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1965 |
have [simp]: "norm (y - a) \<le> r" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1966 |
using that by (simp add: dist_norm norm_minus_commute) |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1967 |
have "norm (g y) \<le> norm (g p)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1968 |
using pge [OF that] by simp |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1969 |
then have "norm (deriv f y) * abs (r - norm (y - a)) \<le> norm (deriv f p) * abs (r - norm (p - a))" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1970 |
by (simp only: dist_norm g_def norm_mult norm_of_real) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1971 |
with that \<open>norm (p - a) < r\<close> show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1972 |
by (simp add: dist_norm divide_simps) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1973 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1974 |
have le_norm_dfp: "r / (r - norm (p - a)) \<le> norm (deriv f p)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1975 |
using gen_le_dfp [of a] \<open>r > 0\<close> by auto |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1976 |
have 1: "f holomorphic_on cball p t" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1977 |
apply (rule holomorphic_on_subset [OF holf]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1978 |
using cpt \<open>r < 1\<close> order_subst1 subset_ball by auto |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1979 |
have 2: "norm (deriv f z) \<le> 2 * norm (deriv f p)" if "z \<in> ball p t" for z |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1980 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1981 |
have z: "z \<in> cball a r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1982 |
by (meson ball_subset_cball subsetD cpt that) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1983 |
then have "norm(z - a) < r" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1984 |
by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1985 |
have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \<le> norm (deriv f p)" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1986 |
using gen_le_dfp [OF z] by simp |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1987 |
with \<open>norm (z - a) < r\<close> \<open>norm (p - a) < r\<close> |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1988 |
have "norm (deriv f z) \<le> (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1989 |
by (simp add: field_simps) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1990 |
also have "... \<le> 2 * norm (deriv f p)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1991 |
apply (rule mult_right_mono) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
1992 |
using that \<open>norm (p - a) < r\<close> \<open>norm(z - a) < r\<close> |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1993 |
apply (simp_all add: field_simps t_def dist_norm [symmetric]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1994 |
using dist_triangle3 [of z a p] by linarith |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1995 |
finally show ?thesis . |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1996 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1997 |
have sqrt2: "sqrt 2 < 2113/1494" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1998 |
by (rule real_less_lsqrt) (auto simp: power2_eq_square) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
1999 |
then have sq3: "0 < 3 - 2 * sqrt 2" by simp |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2000 |
have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2001 |
using sq3 sqrt2 by (auto simp: field_simps r_def) |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2002 |
also have "... \<le> cmod (deriv f p) * (r - cmod (p - a))" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2003 |
using \<open>norm (p - a) < r\<close> le_norm_dfp by (simp add: pos_divide_le_eq) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2004 |
finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2005 |
using pos_divide_less_eq half_gt_zero_iff sq3 by blast |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2006 |
then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2007 |
using sq3 by (simp add: mult.commute t_def) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2008 |
have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball p t" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2009 |
by (rule Bloch_lemma [OF 1 \<open>0 < t\<close> 2]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2010 |
also have "... \<subseteq> f ` ball a 1" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2011 |
apply (rule image_mono) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2012 |
apply (rule order_trans [OF ball_subset_cball]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2013 |
apply (rule order_trans [OF cpt]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2014 |
using \<open>0 < t\<close> \<open>r < 1\<close> apply (simp add: ball_subset_ball_iff dist_norm) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2015 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2016 |
finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball a 1" . |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2017 |
with ** show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2018 |
by (rule that) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2019 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2020 |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2021 |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2022 |
theorem Bloch: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2023 |
assumes holf: "f holomorphic_on ball a r" and "0 < r" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2024 |
and r': "r' \<le> r * norm (deriv f a) / 12" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2025 |
obtains b where "ball b r' \<subseteq> f ` (ball a r)" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2026 |
proof (cases "deriv f a = 0") |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2027 |
case True with r' show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2028 |
using ball_eq_empty that by fastforce |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2029 |
next |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2030 |
case False |
63040 | 2031 |
define C where "C = deriv f a" |
2032 |
have "0 < norm C" using False by (simp add: C_def) |
|
2033 |
have dfa: "f field_differentiable at a" |
|
2034 |
apply (rule holomorphic_on_imp_differentiable_at [OF holf]) |
|
2035 |
using \<open>0 < r\<close> by auto |
|
2036 |
have fo: "(\<lambda>z. f (a + of_real r * z)) = f o (\<lambda>z. (a + of_real r * z))" |
|
2037 |
by (simp add: o_def) |
|
2038 |
have holf': "f holomorphic_on (\<lambda>z. a + complex_of_real r * z) ` ball 0 1" |
|
2039 |
apply (rule holomorphic_on_subset [OF holf]) |
|
2040 |
using \<open>0 < r\<close> apply (force simp: dist_norm norm_mult) |
|
2041 |
done |
|
2042 |
have 1: "(\<lambda>z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" |
|
2043 |
apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ |
|
2044 |
using \<open>0 < r\<close> by (simp add: C_def False) |
|
2045 |
have "((\<lambda>z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2046 |
(deriv f (a + of_real r * z) / C)) (at z)" |
63040 | 2047 |
if "norm z < 1" for z |
2048 |
proof - |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2049 |
have *: "((\<lambda>x. f (a + of_real r * x)) has_field_derivative |
63040 | 2050 |
(deriv f (a + of_real r * z) * of_real r)) (at z)" |
2051 |
apply (simp add: fo) |
|
62534
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
2052 |
apply (rule DERIV_chain [OF field_differentiable_derivI]) |
63040 | 2053 |
apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) |
2054 |
using \<open>0 < r\<close> apply (simp add: dist_norm norm_mult that) |
|
2055 |
apply (rule derivative_eq_intros | simp)+ |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2056 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2057 |
show ?thesis |
63040 | 2058 |
apply (rule derivative_eq_intros * | simp)+ |
2059 |
using \<open>0 < r\<close> by (auto simp: C_def False) |
|
2060 |
qed |
|
2061 |
have 2: "deriv (\<lambda>z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" |
|
2062 |
apply (subst deriv_cdivide_right) |
|
2063 |
apply (simp add: field_differentiable_def fo) |
|
2064 |
apply (rule exI) |
|
2065 |
apply (rule DERIV_chain [OF field_differentiable_derivI]) |
|
2066 |
apply (simp add: dfa) |
|
2067 |
apply (rule derivative_eq_intros | simp add: C_def False fo)+ |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2068 |
using \<open>0 < r\<close> |
63040 | 2069 |
apply (simp add: C_def False fo) |
2070 |
apply (simp add: derivative_intros dfa complex_derivative_chain) |
|
2071 |
done |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2072 |
have sb1: "op * (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1 |
63040 | 2073 |
\<subseteq> f ` ball a r" |
2074 |
using \<open>0 < r\<close> by (auto simp: dist_norm norm_mult C_def False) |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2075 |
have sb2: "ball (C * r * b) r' \<subseteq> op * (C * r) ` ball b t" |
63040 | 2076 |
if "1 / 12 < t" for b t |
2077 |
proof - |
|
2078 |
have *: "r * cmod (deriv f a) / 12 \<le> r * (t * cmod (deriv f a))" |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2079 |
using that \<open>0 < r\<close> less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right |
63040 | 2080 |
by auto |
2081 |
show ?thesis |
|
2082 |
apply clarify |
|
2083 |
apply (rule_tac x="x / (C * r)" in image_eqI) |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2084 |
using \<open>0 < r\<close> |
63040 | 2085 |
apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) |
2086 |
apply (erule less_le_trans) |
|
2087 |
apply (rule order_trans [OF r' *]) |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2088 |
done |
63040 | 2089 |
qed |
2090 |
show ?thesis |
|
2091 |
apply (rule Bloch_unit [OF 1 2]) |
|
2092 |
apply (rename_tac t) |
|
2093 |
apply (rule_tac b="(C * of_real r) * b" in that) |
|
2094 |
apply (drule image_mono [where f = "\<lambda>z. (C * of_real r) * z"]) |
|
2095 |
using sb1 sb2 |
|
2096 |
apply force |
|
2097 |
done |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2098 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2099 |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2100 |
corollary Bloch_general: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2101 |
assumes holf: "f holomorphic_on s" and "a \<in> s" |
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2102 |
and tle: "\<And>z. z \<in> frontier s \<Longrightarrow> t \<le> dist a z" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2103 |
and rle: "r \<le> t * norm(deriv f a) / 12" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2104 |
obtains b where "ball b r \<subseteq> f ` s" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2105 |
proof - |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2106 |
consider "r \<le> 0" | "0 < t * norm(deriv f a) / 12" using rle by force |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2107 |
then show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2108 |
proof cases |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2109 |
case 1 then show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2110 |
by (simp add: Topology_Euclidean_Space.ball_empty that) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2111 |
next |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2112 |
case 2 |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2113 |
show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2114 |
proof (cases "deriv f a = 0") |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2115 |
case True then show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2116 |
using rle by (simp add: Topology_Euclidean_Space.ball_empty that) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2117 |
next |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2118 |
case False |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2119 |
then have "t > 0" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2120 |
using 2 by (force simp: zero_less_mult_iff) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2121 |
have "~ ball a t \<subseteq> s \<Longrightarrow> ball a t \<inter> frontier s \<noteq> {}" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2122 |
apply (rule connected_Int_frontier [of "ball a t" s], simp_all) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2123 |
using \<open>0 < t\<close> \<open>a \<in> s\<close> centre_in_ball apply blast |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2124 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2125 |
with tle have *: "ball a t \<subseteq> s" by fastforce |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2126 |
then have 1: "f holomorphic_on ball a t" |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2127 |
using holf using holomorphic_on_subset by blast |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2128 |
show ?thesis |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2129 |
apply (rule Bloch [OF 1 \<open>t > 0\<close> rle]) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2130 |
apply (rule_tac b=b in that) |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2131 |
using * apply force |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2132 |
done |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2133 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2134 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2135 |
qed |
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62463
diff
changeset
|
2136 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2137 |
subsection \<open>Foundations of Cauchy's residue theorem\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2138 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2139 |
text\<open>Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2140 |
Interactive Theorem Proving\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2141 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2142 |
definition residue :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2143 |
"residue f z = (SOME int. \<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e |
63589 | 2144 |
\<longrightarrow> (f has_contour_integral 2*pi* \<i> *int) (circlepath z \<epsilon>))" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2145 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2146 |
lemma contour_integral_circlepath_eq: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2147 |
assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0<e1" "e1\<le>e2" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2148 |
and e2_cball:"cball z e2 \<subseteq> s" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2149 |
shows |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2150 |
"f contour_integrable_on circlepath z e1" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2151 |
"f contour_integrable_on circlepath z e2" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2152 |
"contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2153 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2154 |
define l where "l \<equiv> linepath (z+e2) (z+e1)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2155 |
have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2156 |
have "e2>0" using \<open>e1>0\<close> \<open>e1\<le>e2\<close> by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2157 |
have zl_img:"z\<notin>path_image l" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2158 |
proof |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2159 |
assume "z \<in> path_image l" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2160 |
then have "e2 \<le> cmod (e2 - e1)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2161 |
using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \<open>e1>0\<close> \<open>e2>0\<close> unfolding l_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2162 |
by (auto simp add:closed_segment_commute) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2163 |
thus False using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2164 |
apply (subst (asm) norm_of_real) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2165 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2166 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2167 |
define g where "g \<equiv> circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2168 |
show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2169 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2170 |
show "f contour_integrable_on circlepath z e2" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2171 |
apply (intro contour_integrable_continuous_circlepath[OF |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2172 |
continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2173 |
using \<open>e2>0\<close> e2_cball by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2174 |
show "f contour_integrable_on (circlepath z e1)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2175 |
apply (intro contour_integrable_continuous_circlepath[OF |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2176 |
continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2177 |
using \<open>e1>0\<close> \<open>e1\<le>e2\<close> e2_cball by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2178 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2179 |
have [simp]:"f contour_integrable_on l" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2180 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2181 |
have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2182 |
by (intro closed_segment_subset,auto simp add:dist_norm) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2183 |
hence "closed_segment (z + e2) (z + e1) \<subseteq> s - {z}" using zl_img e2_cball unfolding l_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2184 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2185 |
then show "f contour_integrable_on l" unfolding l_def |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2186 |
apply (intro contour_integrable_continuous_linepath[OF |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2187 |
continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2188 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2189 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2190 |
let ?ig="\<lambda>g. contour_integral g f" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2191 |
have "(f has_contour_integral 0) g" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2192 |
proof (rule Cauchy_theorem_global[OF _ f_holo]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2193 |
show "open (s - {z})" using \<open>open s\<close> by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2194 |
show "valid_path g" unfolding g_def l_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2195 |
show "pathfinish g = pathstart g" unfolding g_def l_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2196 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2197 |
have path_img:"path_image g \<subseteq> cball z e2" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2198 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2199 |
have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close> |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2200 |
by (intro closed_segment_subset,auto simp add:dist_norm) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2201 |
moreover have "sphere z \<bar>e1\<bar> \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1\<le>e2\<close> \<open>e1>0\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2202 |
ultimately show ?thesis unfolding g_def l_def using \<open>e2>0\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2203 |
by (simp add: path_image_join closed_segment_commute) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2204 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2205 |
show "path_image g \<subseteq> s - {z}" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2206 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2207 |
have "z\<notin>path_image g" using zl_img |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2208 |
unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2209 |
moreover note \<open>cball z e2 \<subseteq> s\<close> and path_img |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2210 |
ultimately show ?thesis by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2211 |
qed |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2212 |
show "winding_number g w = 0" when"w \<notin> s - {z}" for w |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2213 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2214 |
have "winding_number g w = 0" when "w\<notin>s" using that e2_cball |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2215 |
apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2216 |
by (auto simp add:g_def l_def) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2217 |
moreover have "winding_number g z=0" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2218 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2219 |
let ?Wz="\<lambda>g. winding_number g z" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2220 |
have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2221 |
+ ?Wz (reversepath l)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2222 |
using \<open>e2>0\<close> \<open>e1>0\<close> zl_img unfolding g_def l_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2223 |
by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2224 |
also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2225 |
using zl_img |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2226 |
apply (subst (2) winding_number_reversepath) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2227 |
by (auto simp add:l_def closed_segment_commute) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2228 |
also have "... = 0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2229 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2230 |
have "?Wz (circlepath z e2) = 1" using \<open>e2>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2231 |
by (auto intro: winding_number_circlepath_centre) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2232 |
moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \<open>e1>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2233 |
apply (subst winding_number_reversepath) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2234 |
by (auto intro: winding_number_circlepath_centre) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2235 |
ultimately show ?thesis by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2236 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2237 |
finally show ?thesis . |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2238 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2239 |
ultimately show ?thesis using that by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2240 |
qed |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2241 |
qed |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2242 |
then have "0 = ?ig g" using contour_integral_unique by simp |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2243 |
also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2244 |
+ ?ig (reversepath l)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2245 |
unfolding g_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2246 |
by (auto simp add:contour_integrable_reversepath_eq) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2247 |
also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2248 |
by (auto simp add:contour_integral_reversepath) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2249 |
finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2250 |
by simp |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2251 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2252 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2253 |
lemma base_residue: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2254 |
assumes "open s" "z\<in>s" "r>0" and f_holo:"f holomorphic_on (s - {z})" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2255 |
and r_cball:"cball z r \<subseteq> s" |
63589 | 2256 |
shows "(f has_contour_integral 2 * pi * \<i> * (residue f z)) (circlepath z r)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2257 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2258 |
obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2259 |
using open_contains_cball[of s] \<open>open s\<close> \<open>z\<in>s\<close> by auto |
63589 | 2260 |
define c where "c \<equiv> 2 * pi * \<i>" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2261 |
define i where "i \<equiv> contour_integral (circlepath z e) f / c" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2262 |
have "(f has_contour_integral c*i) (circlepath z \<epsilon>)" when "\<epsilon>>0" "\<epsilon><e" for \<epsilon> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2263 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2264 |
have "contour_integral (circlepath z e) f = contour_integral (circlepath z \<epsilon>) f" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2265 |
"f contour_integrable_on circlepath z \<epsilon>" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2266 |
"f contour_integrable_on circlepath z e" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2267 |
using \<open>\<epsilon><e\<close> |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2268 |
by (intro contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> _ e_cball],auto)+ |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2269 |
then show ?thesis unfolding i_def c_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2270 |
by (auto intro:has_contour_integral_integral) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2271 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2272 |
then have "\<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2273 |
unfolding residue_def c_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2274 |
apply (rule_tac someI[of _ i],intro exI[where x=e]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2275 |
by (auto simp add:\<open>e>0\<close> c_def) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2276 |
then obtain e' where "e'>0" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2277 |
and e'_def:"\<forall>\<epsilon>>0. \<epsilon><e' \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2278 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2279 |
let ?int="\<lambda>e. contour_integral (circlepath z e) f" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2280 |
def \<epsilon>\<equiv>"Min {r,e'} / 2" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2281 |
have "\<epsilon>>0" "\<epsilon>\<le>r" "\<epsilon><e'" using \<open>r>0\<close> \<open>e'>0\<close> unfolding \<epsilon>_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2282 |
have "(f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2283 |
using e'_def[rule_format,OF \<open>\<epsilon>>0\<close> \<open>\<epsilon><e'\<close>] . |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2284 |
then show ?thesis unfolding c_def |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2285 |
using contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> \<open>\<epsilon>\<le>r\<close> r_cball] |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2286 |
by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \<epsilon>" "circlepath z r"]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2287 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2288 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2289 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2290 |
lemma residue_holo: |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2291 |
assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2292 |
shows "residue f z = 0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2293 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2294 |
define c where "c \<equiv> 2 * pi * \<i>" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2295 |
obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2296 |
using open_contains_cball_eq by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2297 |
have "(f has_contour_integral c*residue f z) (circlepath z e)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2298 |
using f_holo |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2299 |
by (auto intro: base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def]) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2300 |
moreover have "(f has_contour_integral 0) (circlepath z e)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2301 |
using f_holo e_cball \<open>e>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2302 |
by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2303 |
ultimately have "c*residue f z =0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2304 |
using has_contour_integral_unique by blast |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2305 |
thus ?thesis unfolding c_def by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2306 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2307 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2308 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2309 |
lemma residue_const:"residue (\<lambda>_. c) z = 0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2310 |
by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2311 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2312 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2313 |
lemma residue_add: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2314 |
assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2315 |
and g_holo:"g holomorphic_on s - {z}" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2316 |
shows "residue (\<lambda>z. f z + g z) z= residue f z + residue g z" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2317 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2318 |
define c where "c \<equiv> 2 * pi * \<i>" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2319 |
define fg where "fg \<equiv> (\<lambda>z. f z+g z)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2320 |
obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2321 |
using open_contains_cball_eq by blast |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2322 |
have "(fg has_contour_integral c * residue fg z) (circlepath z e)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2323 |
unfolding fg_def using f_holo g_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2324 |
apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2325 |
by (auto intro:holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2326 |
moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2327 |
unfolding fg_def using f_holo g_holo |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2328 |
by (auto intro: has_contour_integral_add base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def]) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2329 |
ultimately have "c*(residue f z + residue g z) = c * residue fg z" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2330 |
using has_contour_integral_unique by (auto simp add:distrib_left) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2331 |
thus ?thesis unfolding fg_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2332 |
by (auto simp add:c_def) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2333 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2334 |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2335 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2336 |
lemma residue_lmul: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2337 |
assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2338 |
shows "residue (\<lambda>z. c * (f z)) z= c * residue f z" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2339 |
proof (cases "c=0") |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2340 |
case True |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2341 |
thus ?thesis using residue_const by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2342 |
next |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2343 |
case False |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2344 |
def c'\<equiv>"2 * pi * \<i>" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2345 |
def f'\<equiv>"(\<lambda>z. c * (f z))" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2346 |
obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2347 |
using open_contains_cball_eq by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2348 |
have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2349 |
unfolding f'_def using f_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2350 |
apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2351 |
by (auto intro:holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2352 |
moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2353 |
unfolding f'_def using f_holo |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2354 |
by (auto intro: has_contour_integral_lmul |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2355 |
base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def]) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2356 |
ultimately have "c' * residue f' z = c * (c' * residue f z)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2357 |
using has_contour_integral_unique by auto |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2358 |
thus ?thesis unfolding f'_def c'_def using False |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2359 |
by (auto simp add:field_simps) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2360 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2361 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2362 |
lemma residue_rmul: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2363 |
assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2364 |
shows "residue (\<lambda>z. (f z) * c) z= residue f z * c" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2365 |
using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2366 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2367 |
lemma residue_div: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2368 |
assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2369 |
shows "residue (\<lambda>z. (f z) / c) z= residue f z / c " |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2370 |
using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2371 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2372 |
lemma residue_neg: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2373 |
assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2374 |
shows "residue (\<lambda>z. - (f z)) z= - residue f z" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2375 |
using residue_lmul[OF assms,of "-1"] by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2376 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2377 |
lemma residue_diff: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2378 |
assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2379 |
and g_holo:"g holomorphic_on s - {z}" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2380 |
shows "residue (\<lambda>z. f z - g z) z= residue f z - residue g z" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2381 |
using residue_add[OF assms(1,2,3),of "\<lambda>z. - g z"] residue_neg[OF assms(1,2,4)] |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2382 |
by (auto intro:holomorphic_intros g_holo) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2383 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2384 |
lemma residue_simple: |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2385 |
assumes "open s" "z\<in>s" and f_holo:"f holomorphic_on s" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2386 |
shows "residue (\<lambda>w. f w / (w - z)) z = f z" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2387 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2388 |
define c where "c \<equiv> 2 * pi * \<i>" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2389 |
def f'\<equiv>"\<lambda>w. f w / (w - z)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2390 |
obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2391 |
using open_contains_cball_eq by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2392 |
have "(f' has_contour_integral c * f z) (circlepath z e)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2393 |
unfolding f'_def c_def using \<open>e>0\<close> f_holo e_cball |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2394 |
by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2395 |
moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2396 |
unfolding f'_def using f_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2397 |
apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2398 |
by (auto intro!:holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2399 |
ultimately have "c * f z = c * residue f' z" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2400 |
using has_contour_integral_unique by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2401 |
thus ?thesis unfolding c_def f'_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2402 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2403 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2404 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2405 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2406 |
subsubsection \<open>Cauchy's residue theorem\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2407 |
|
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2408 |
lemma get_integrable_path: |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2409 |
assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\<in>s-pts" "b\<in>s-pts" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2410 |
obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2411 |
"path_image g \<subseteq> s-pts" "f contour_integrable_on g" using assms |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2412 |
proof (induct arbitrary:s thesis a rule:finite_induct[OF \<open>finite pts\<close>]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2413 |
case 1 |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2414 |
obtain g where "valid_path g" "path_image g \<subseteq> s" "pathstart g = a" "pathfinish g = b" |
62837 | 2415 |
using connected_open_polynomial_connected[OF \<open>open s\<close>,of a b ] \<open>connected (s - {})\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2416 |
valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2417 |
moreover have "f contour_integrable_on g" |
62837 | 2418 |
using contour_integrable_holomorphic_simple[OF _ \<open>open s\<close> \<open>valid_path g\<close> \<open>path_image g \<subseteq> s\<close>,of f] |
2419 |
\<open>f holomorphic_on s - {}\<close> |
|
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2420 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2421 |
ultimately show ?case using "1"(1)[of g] by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2422 |
next |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2423 |
case idt:(2 p pts) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2424 |
obtain e where "e>0" and e:"\<forall>w\<in>ball a e. w \<in> s \<and> (w \<noteq> a \<longrightarrow> w \<notin> insert p pts)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2425 |
using finite_ball_avoid[OF \<open>open s\<close> \<open>finite (insert p pts)\<close>, of a] |
62837 | 2426 |
\<open>a \<in> s - insert p pts\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2427 |
by auto |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2428 |
define a' where "a' \<equiv> a+e/2" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2429 |
have "a'\<in>s-{p} -pts" using e[rule_format,of "a+e/2"] \<open>e>0\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2430 |
by (auto simp add:dist_complex_def a'_def) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2431 |
then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2432 |
"path_image g' \<subseteq> s - {p} - pts" "f contour_integrable_on g'" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2433 |
using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2434 |
by (metis Diff_insert2 open_delete) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2435 |
define g where "g \<equiv> linepath a a' +++ g'" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2436 |
have "valid_path g" unfolding g_def by (auto intro: valid_path_join) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2437 |
moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2438 |
moreover have "path_image g \<subseteq> s - insert p pts" unfolding g_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2439 |
proof (rule subset_path_image_join) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2440 |
have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2441 |
by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2442 |
then show "path_image (linepath a a') \<subseteq> s - insert p pts" using e idt(9) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2443 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2444 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2445 |
show "path_image g' \<subseteq> s - insert p pts" using g'(4) by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2446 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2447 |
moreover have "f contour_integrable_on g" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2448 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2449 |
have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2450 |
by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2451 |
then have "continuous_on (closed_segment a a') f" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2452 |
using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2453 |
apply (elim continuous_on_subset) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2454 |
by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2455 |
then have "f contour_integrable_on linepath a a'" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2456 |
using contour_integrable_continuous_linepath by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2457 |
then show ?thesis unfolding g_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2458 |
apply (rule contour_integrable_joinI) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2459 |
by (auto simp add: \<open>e>0\<close>) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2460 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2461 |
ultimately show ?case using idt.prems(1)[of g] by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2462 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2463 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2464 |
lemma Cauchy_theorem_aux: |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2465 |
assumes "open s" "connected (s-pts)" "finite pts" "pts \<subseteq> s" "f holomorphic_on s-pts" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2466 |
"valid_path g" "pathfinish g = pathstart g" "path_image g \<subseteq> s-pts" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2467 |
"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2468 |
"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2469 |
shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2470 |
using assms |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2471 |
proof (induct arbitrary:s g rule:finite_induct[OF \<open>finite pts\<close>]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2472 |
case 1 |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2473 |
then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2474 |
next |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2475 |
case (2 p pts) |
62837 | 2476 |
note fin[simp] = \<open>finite (insert p pts)\<close> |
2477 |
and connected = \<open>connected (s - insert p pts)\<close> |
|
2478 |
and valid[simp] = \<open>valid_path g\<close> |
|
2479 |
and g_loop[simp] = \<open>pathfinish g = pathstart g\<close> |
|
2480 |
and holo[simp]= \<open>f holomorphic_on s - insert p pts\<close> |
|
2481 |
and path_img = \<open>path_image g \<subseteq> s - insert p pts\<close> |
|
2482 |
and winding = \<open>\<forall>z. z \<notin> s \<longrightarrow> winding_number g z = 0\<close> |
|
2483 |
and h = \<open>\<forall>pa\<in>s. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s \<and> (w \<noteq> pa \<longrightarrow> w \<notin> insert p pts))\<close> |
|
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2484 |
have "h p>0" and "p\<in>s" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2485 |
and h_p: "\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> insert p pts)" |
62837 | 2486 |
using h \<open>insert p pts \<subseteq> s\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2487 |
obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2488 |
"path_image pg \<subseteq> s-insert p pts" "f contour_integrable_on pg" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2489 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2490 |
have "p + h p\<in>cball p (h p)" using h[rule_format,of p] |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2491 |
by (simp add: \<open>p \<in> s\<close> dist_norm) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2492 |
then have "p + h p \<in> s - insert p pts" using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2493 |
by fastforce |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2494 |
moreover have "pathstart g \<in> s - insert p pts " using path_img by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2495 |
ultimately show ?thesis |
62837 | 2496 |
using get_integrable_path[OF \<open>open s\<close> connected fin holo,of "pathstart g" "p+h p"] that |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2497 |
by blast |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2498 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2499 |
obtain n::int where "n=winding_number g p" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2500 |
using integer_winding_number[OF _ g_loop,of p] valid path_img |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2501 |
by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2502 |
define p_circ where "p_circ \<equiv> circlepath p (h p)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2503 |
define p_circ_pt where "p_circ_pt \<equiv> linepath (p+h p) (p+h p)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2504 |
define n_circ where "n_circ \<equiv> \<lambda>n. (op +++ p_circ ^^ n) p_circ_pt" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2505 |
define cp where "cp \<equiv> if n\<ge>0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2506 |
have n_circ:"valid_path (n_circ k)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2507 |
"winding_number (n_circ k) p = k" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2508 |
"pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2509 |
"path_image (n_circ k) = (if k=0 then {p + h p} else sphere p (h p))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2510 |
"p \<notin> path_image (n_circ k)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2511 |
"\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number (n_circ k) p'=0 \<and> p'\<notin>path_image (n_circ k)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2512 |
"f contour_integrable_on (n_circ k)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2513 |
"contour_integral (n_circ k) f = k * contour_integral p_circ f" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2514 |
for k |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2515 |
proof (induct k) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2516 |
case 0 |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2517 |
show "valid_path (n_circ 0)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2518 |
and "path_image (n_circ 0) = (if 0=0 then {p + h p} else sphere p (h p))" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2519 |
and "winding_number (n_circ 0) p = of_nat 0" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2520 |
and "pathstart (n_circ 0) = p + h p" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2521 |
and "pathfinish (n_circ 0) = p + h p" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2522 |
and "p \<notin> path_image (n_circ 0)" |
62837 | 2523 |
unfolding n_circ_def p_circ_pt_def using \<open>h p > 0\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2524 |
by (auto simp add: dist_norm) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2525 |
show "winding_number (n_circ 0) p'=0 \<and> p'\<notin>path_image (n_circ 0)" when "p'\<notin>s- pts" for p' |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2526 |
unfolding n_circ_def p_circ_pt_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2527 |
apply (auto intro!:winding_number_trivial) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2528 |
by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+ |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2529 |
show "f contour_integrable_on (n_circ 0)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2530 |
unfolding n_circ_def p_circ_pt_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2531 |
by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2532 |
show "contour_integral (n_circ 0) f = of_nat 0 * contour_integral p_circ f" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2533 |
unfolding n_circ_def p_circ_pt_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2534 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2535 |
case (Suc k) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2536 |
have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2537 |
have pcirc:"p \<notin> path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" |
62837 | 2538 |
using Suc(3) unfolding p_circ_def using \<open>h p > 0\<close> by (auto simp add: p_circ_def) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2539 |
have pcirc_image:"path_image p_circ \<subseteq> s - insert p pts" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2540 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2541 |
have "path_image p_circ \<subseteq> cball p (h p)" using \<open>0 < h p\<close> p_circ_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2542 |
then show ?thesis using h_p pcirc(1) by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2543 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2544 |
have pcirc_integrable:"f contour_integrable_on p_circ" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2545 |
by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2546 |
contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2547 |
holomorphic_on_subset[OF holo]) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2548 |
show "valid_path (n_circ (Suc k))" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2549 |
using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2550 |
show "path_image (n_circ (Suc k)) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2551 |
= (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2552 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2553 |
have "path_image p_circ = sphere p (h p)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2554 |
unfolding p_circ_def using \<open>0 < h p\<close> by auto |
62837 | 2555 |
then show ?thesis unfolding n_Suc using Suc.hyps(5) \<open>h p>0\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2556 |
by (auto simp add: path_image_join[OF pcirc(3)] dist_norm) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2557 |
qed |
62837 | 2558 |
then show "p \<notin> path_image (n_circ (Suc k))" using \<open>h p>0\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2559 |
show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2560 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2561 |
have "winding_number p_circ p = 1" |
62837 | 2562 |
by (simp add: \<open>h p > 0\<close> p_circ_def winding_number_circlepath_centre) |
2563 |
moreover have "p \<notin> path_image (n_circ k)" using Suc(5) \<open>h p>0\<close> by auto |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2564 |
then have "winding_number (p_circ +++ n_circ k) p |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2565 |
= winding_number p_circ p + winding_number (n_circ k) p" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2566 |
using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2567 |
apply (intro winding_number_join) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2568 |
by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2569 |
ultimately show ?thesis using Suc(2) unfolding n_circ_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2570 |
by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2571 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2572 |
show "pathstart (n_circ (Suc k)) = p + h p" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2573 |
by (simp add: n_circ_def p_circ_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2574 |
show "pathfinish (n_circ (Suc k)) = p + h p" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2575 |
using Suc(4) unfolding n_circ_def by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2576 |
show "winding_number (n_circ (Suc k)) p'=0 \<and> p'\<notin>path_image (n_circ (Suc k))" when "p'\<notin>s-pts" for p' |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2577 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2578 |
have " p' \<notin> path_image p_circ" using \<open>p \<in> s\<close> h p_circ_def that using pcirc_image by blast |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2579 |
moreover have "p' \<notin> path_image (n_circ k)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2580 |
using Suc.hyps(7) that by blast |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2581 |
moreover have "winding_number p_circ p' = 0" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2582 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2583 |
have "path_image p_circ \<subseteq> cball p (h p)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2584 |
using h unfolding p_circ_def using \<open>p \<in> s\<close> by fastforce |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2585 |
moreover have "p'\<notin>cball p (h p)" using \<open>p \<in> s\<close> h that "2.hyps"(2) by fastforce |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2586 |
ultimately show ?thesis unfolding p_circ_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2587 |
apply (intro winding_number_zero_outside) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2588 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2589 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2590 |
ultimately show ?thesis |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2591 |
unfolding n_Suc |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2592 |
apply (subst winding_number_join) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2593 |
by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2594 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2595 |
show "f contour_integrable_on (n_circ (Suc k))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2596 |
unfolding n_Suc |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2597 |
by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2598 |
show "contour_integral (n_circ (Suc k)) f = (Suc k) * contour_integral p_circ f" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2599 |
unfolding n_Suc |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2600 |
by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2601 |
Suc(9) algebra_simps) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2602 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2603 |
have cp[simp]:"pathstart cp = p + h p" "pathfinish cp = p + h p" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2604 |
"valid_path cp" "path_image cp \<subseteq> s - insert p pts" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2605 |
"winding_number cp p = - n" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2606 |
"\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number cp p'=0 \<and> p' \<notin> path_image cp" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2607 |
"f contour_integrable_on cp" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2608 |
"contour_integral cp f = - n * contour_integral p_circ f" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2609 |
proof - |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2610 |
show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2611 |
using n_circ unfolding cp_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2612 |
next |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2613 |
have "sphere p (h p) \<subseteq> s - insert p pts" |
62837 | 2614 |
using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> by force |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2615 |
moreover have "p + complex_of_real (h p) \<in> s - insert p pts" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2616 |
using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2617 |
ultimately show "path_image cp \<subseteq> s - insert p pts" unfolding cp_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2618 |
using n_circ(5) by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2619 |
next |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2620 |
show "winding_number cp p = - n" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2621 |
unfolding cp_def using winding_number_reversepath n_circ \<open>h p>0\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2622 |
by (auto simp: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2623 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2624 |
show "winding_number cp p'=0 \<and> p' \<notin> path_image cp" when "p'\<notin>s - pts" for p' |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2625 |
unfolding cp_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2626 |
apply (auto) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2627 |
apply (subst winding_number_reversepath) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2628 |
by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1)) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2629 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2630 |
show "f contour_integrable_on cp" unfolding cp_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2631 |
using contour_integrable_reversepath_eq n_circ(1,8) by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2632 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2633 |
show "contour_integral cp f = - n * contour_integral p_circ f" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2634 |
unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2635 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2636 |
qed |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2637 |
def g'\<equiv>"g +++ pg +++ cp +++ (reversepath pg)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2638 |
have "contour_integral g' f = (\<Sum>p\<in>pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2639 |
proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \<open>finite pts\<close> ]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2640 |
show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) |
62837 | 2641 |
show "open (s - {p})" using \<open>open s\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2642 |
show " pts \<subseteq> s - {p}" using \<open>insert p pts \<subseteq> s\<close> \<open> p \<notin> pts\<close> by blast |
62837 | 2643 |
show "f holomorphic_on s - {p} - pts" using holo \<open>p \<notin> pts\<close> by (metis Diff_insert2) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2644 |
show "valid_path g'" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2645 |
unfolding g'_def cp_def using n_circ valid pg g_loop |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2646 |
by (auto intro!:valid_path_join ) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2647 |
show "pathfinish g' = pathstart g'" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2648 |
unfolding g'_def cp_def using pg(2) by simp |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2649 |
show "path_image g' \<subseteq> s - {p} - pts" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2650 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2651 |
def s'\<equiv>"s - {p} - pts" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2652 |
have s':"s' = s-insert p pts " unfolding s'_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2653 |
then show ?thesis using path_img pg(4) cp(4) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2654 |
unfolding g'_def |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2655 |
apply (fold s'_def s') |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2656 |
apply (intro subset_path_image_join) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2657 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2658 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2659 |
note path_join_imp[simp] |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2660 |
show "\<forall>z. z \<notin> s - {p} \<longrightarrow> winding_number g' z = 0" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2661 |
proof clarify |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2662 |
fix z assume z:"z\<notin>s - {p}" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2663 |
have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2664 |
+ winding_number (pg +++ cp +++ (reversepath pg)) z" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2665 |
proof (rule winding_number_join) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2666 |
show "path g" using \<open>valid_path g\<close> by (simp add: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2667 |
show "z \<notin> path_image g" using z path_img by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2668 |
show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2669 |
by (simp add: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2670 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2671 |
have "path_image (pg +++ cp +++ reversepath pg) \<subseteq> s - insert p pts" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2672 |
using pg(4) cp(4) by (auto simp:subset_path_image_join) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2673 |
then show "z \<notin> path_image (pg +++ cp +++ reversepath pg)" using z by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2674 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2675 |
show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2676 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2677 |
also have "... = winding_number g z + (winding_number pg z |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2678 |
+ winding_number (cp +++ (reversepath pg)) z)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2679 |
proof (subst add_left_cancel,rule winding_number_join) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2680 |
show "path pg" and "path (cp +++ reversepath pg)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2681 |
and "pathfinish pg = pathstart (cp +++ reversepath pg)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2682 |
by (auto simp add: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2683 |
show "z \<notin> path_image pg" using pg(4) z by blast |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2684 |
show "z \<notin> path_image (cp +++ reversepath pg)" using z |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2685 |
by (metis Diff_iff \<open>z \<notin> path_image pg\<close> contra_subsetD cp(4) insertI1 |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2686 |
not_in_path_image_join path_image_reversepath singletonD) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2687 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2688 |
also have "... = winding_number g z + (winding_number pg z |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2689 |
+ (winding_number cp z + winding_number (reversepath pg) z))" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2690 |
apply (auto intro!:winding_number_join simp: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2691 |
apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2692 |
by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2693 |
also have "... = winding_number g z + winding_number cp z" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2694 |
apply (subst winding_number_reversepath) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2695 |
apply (auto simp: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2696 |
by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2697 |
finally have "winding_number g' z = winding_number g z + winding_number cp z" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2698 |
unfolding g'_def . |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2699 |
moreover have "winding_number g z + winding_number cp z = 0" |
62837 | 2700 |
using winding z \<open>n=winding_number g p\<close> by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2701 |
ultimately show "winding_number g' z = 0" unfolding g'_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2702 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2703 |
show "\<forall>pa\<in>s - {p}. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s - {p} \<and> (w \<noteq> pa \<longrightarrow> w \<notin> pts))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2704 |
using h by fastforce |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2705 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2706 |
moreover have "contour_integral g' f = contour_integral g f |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2707 |
- winding_number g p * contour_integral p_circ f" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2708 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2709 |
have "contour_integral g' f = contour_integral g f |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2710 |
+ contour_integral (pg +++ cp +++ reversepath pg) f" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2711 |
unfolding g'_def |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2712 |
apply (subst contour_integral_join) |
62837 | 2713 |
by (auto simp add:open_Diff[OF \<open>open s\<close>,OF finite_imp_closed[OF fin]] |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2714 |
intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2715 |
contour_integrable_reversepath) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2716 |
also have "... = contour_integral g f + contour_integral pg f |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2717 |
+ contour_integral (cp +++ reversepath pg) f" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2718 |
apply (subst contour_integral_join) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2719 |
by (auto simp add:contour_integrable_reversepath) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2720 |
also have "... = contour_integral g f + contour_integral pg f |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2721 |
+ contour_integral cp f + contour_integral (reversepath pg) f" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2722 |
apply (subst contour_integral_join) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2723 |
by (auto simp add:contour_integrable_reversepath) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2724 |
also have "... = contour_integral g f + contour_integral cp f" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2725 |
using contour_integral_reversepath |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2726 |
by (auto simp add:contour_integrable_reversepath) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2727 |
also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f" |
62837 | 2728 |
using \<open>n=winding_number g p\<close> by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2729 |
finally show ?thesis . |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2730 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2731 |
moreover have "winding_number g' p' = winding_number g p'" when "p'\<in>pts" for p' |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2732 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2733 |
have [simp]: "p' \<notin> path_image g" "p' \<notin> path_image pg" "p'\<notin>path_image cp" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2734 |
using "2.prems"(8) that |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2735 |
apply blast |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2736 |
apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2737 |
by (meson DiffD2 cp(4) set_rev_mp subset_insertI that) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2738 |
have "winding_number g' p' = winding_number g p' |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2739 |
+ winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2740 |
apply (subst winding_number_join) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2741 |
apply (simp_all add: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2742 |
apply (intro not_in_path_image_join) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2743 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2744 |
also have "... = winding_number g p' + winding_number pg p' |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2745 |
+ winding_number (cp +++ reversepath pg) p'" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2746 |
apply (subst winding_number_join) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2747 |
apply (simp_all add: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2748 |
apply (intro not_in_path_image_join) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2749 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2750 |
also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p' |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2751 |
+ winding_number (reversepath pg) p'" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2752 |
apply (subst winding_number_join) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2753 |
by (simp_all add: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2754 |
also have "... = winding_number g p' + winding_number cp p'" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2755 |
apply (subst winding_number_reversepath) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2756 |
by (simp_all add: valid_path_imp_path) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2757 |
also have "... = winding_number g p'" using that by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2758 |
finally show ?thesis . |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2759 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2760 |
ultimately show ?case unfolding p_circ_def |
64267 | 2761 |
apply (subst (asm) sum.cong[OF refl, |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2762 |
of pts _ "\<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"]) |
64267 | 2763 |
by (auto simp add:sum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2764 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2765 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2766 |
lemma Cauchy_theorem_singularities: |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2767 |
assumes "open s" "connected s" "finite pts" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2768 |
holo:"f holomorphic_on s-pts" and |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2769 |
"valid_path g" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2770 |
loop:"pathfinish g = pathstart g" and |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2771 |
"path_image g \<subseteq> s-pts" and |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2772 |
homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2773 |
avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2774 |
shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2775 |
(is "?L=?R") |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2776 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2777 |
define circ where "circ \<equiv> \<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2778 |
define pts1 where "pts1 \<equiv> pts \<inter> s" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2779 |
define pts2 where "pts2 \<equiv> pts - pts1" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2780 |
have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> s={}" "pts1\<subseteq>s" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2781 |
unfolding pts1_def pts2_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2782 |
have "contour_integral g f = (\<Sum>p\<in>pts1. circ p)" unfolding circ_def |
62837 | 2783 |
proof (rule Cauchy_theorem_aux[OF \<open>open s\<close> _ _ \<open>pts1\<subseteq>s\<close> _ \<open>valid_path g\<close> loop _ homo]) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2784 |
have "finite pts1" unfolding pts1_def using \<open>finite pts\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2785 |
then show "connected (s - pts1)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2786 |
using \<open>open s\<close> \<open>connected s\<close> connected_open_delete_finite[of s] by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2787 |
next |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2788 |
show "finite pts1" using \<open>pts = pts1 \<union> pts2\<close> assms(3) by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2789 |
show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2790 |
show "path_image g \<subseteq> s - pts1" using assms(7) pts1_def by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2791 |
show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts1))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2792 |
by (simp add: avoid pts1_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2793 |
qed |
64267 | 2794 |
moreover have "sum circ pts2=0" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2795 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2796 |
have "winding_number g p=0" when "p\<in>pts2" for p |
62837 | 2797 |
using \<open>pts2 \<inter> s={}\<close> that homo[rule_format,of p] by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2798 |
thus ?thesis unfolding circ_def |
64267 | 2799 |
apply (intro sum.neutral) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2800 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2801 |
qed |
64267 | 2802 |
moreover have "?R=sum circ pts1 + sum circ pts2" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2803 |
unfolding circ_def |
64267 | 2804 |
using sum.union_disjoint[OF _ _ \<open>pts1 \<inter> pts2 = {}\<close>] \<open>finite pts\<close> \<open>pts=pts1 \<union> pts2\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2805 |
by blast |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2806 |
ultimately show ?thesis |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2807 |
apply (fold circ_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2808 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2809 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2810 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2811 |
lemma Residue_theorem: |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2812 |
fixes s pts::"complex set" and f::"complex \<Rightarrow> complex" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2813 |
and g::"real \<Rightarrow> complex" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2814 |
assumes "open s" "connected s" "finite pts" and |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2815 |
holo:"f holomorphic_on s-pts" and |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2816 |
"valid_path g" and |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2817 |
loop:"pathfinish g = pathstart g" and |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2818 |
"path_image g \<subseteq> s-pts" and |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2819 |
homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2820 |
shows "contour_integral g f = 2 * pi * \<i> *(\<Sum>p\<in>pts. winding_number g p * residue f p)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2821 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2822 |
define c where "c \<equiv> 2 * pi * \<i>" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2823 |
obtain h where avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2824 |
using finite_cball_avoid[OF \<open>open s\<close> \<open>finite pts\<close>] by metis |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2825 |
have "contour_integral g f |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2826 |
= (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2827 |
using Cauchy_theorem_singularities[OF assms avoid] . |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2828 |
also have "... = (\<Sum>p\<in>pts. c * winding_number g p * residue f p)" |
64267 | 2829 |
proof (intro sum.cong) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2830 |
show "pts = pts" by simp |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2831 |
next |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2832 |
fix x assume "x \<in> pts" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2833 |
show "winding_number g x * contour_integral (circlepath x (h x)) f |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2834 |
= c * winding_number g x * residue f x" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2835 |
proof (cases "x\<in>s") |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2836 |
case False |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2837 |
then have "winding_number g x=0" using homo by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2838 |
thus ?thesis by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2839 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2840 |
case True |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2841 |
have "contour_integral (circlepath x (h x)) f = c* residue f x" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2842 |
using \<open>x\<in>pts\<close> \<open>finite pts\<close> avoid[rule_format,OF True] |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2843 |
apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2844 |
by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \<open>open s\<close> finite_imp_closed]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2845 |
then show ?thesis by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2846 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2847 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2848 |
also have "... = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)" |
64267 | 2849 |
by (simp add: sum_distrib_left algebra_simps) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2850 |
finally show ?thesis unfolding c_def . |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2851 |
qed |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2852 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2853 |
subsection \<open>The argument principle\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2854 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2855 |
definition is_pole :: "('a::topological_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool" where |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2856 |
"is_pole f a = (LIM x (at a). f x :> at_infinity)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2857 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2858 |
lemma is_pole_tendsto: |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2859 |
fixes f::"('a::topological_space \<Rightarrow> 'b::real_normed_div_algebra)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2860 |
shows "is_pole f x \<Longrightarrow> ((inverse o f) \<longlongrightarrow> 0) (at x)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2861 |
unfolding is_pole_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2862 |
by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2863 |
|
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2864 |
lemma is_pole_inverse_holomorphic: |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2865 |
assumes "open s" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2866 |
and f_holo:"f holomorphic_on (s-{z})" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2867 |
and pole:"is_pole f z" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2868 |
and non_z:"\<forall>x\<in>s-{z}. f x\<noteq>0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2869 |
shows "(\<lambda>x. if x=z then 0 else inverse (f x)) holomorphic_on s" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2870 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2871 |
define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2872 |
have "isCont g z" unfolding isCont_def using is_pole_tendsto[OF pole] |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2873 |
apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \<circ> f"]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2874 |
by (simp_all add:g_def) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2875 |
moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2876 |
hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2877 |
by (auto elim!:continuous_on_inverse simp add:non_z) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2878 |
hence "continuous_on (s-{z}) g" unfolding g_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2879 |
apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2880 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2881 |
ultimately have "continuous_on s g" using open_delete[OF \<open>open s\<close>] \<open>open s\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2882 |
by (auto simp add:continuous_on_eq_continuous_at) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2883 |
moreover have "(inverse o f) holomorphic_on (s-{z})" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2884 |
unfolding comp_def using f_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2885 |
by (auto elim!:holomorphic_on_inverse simp add:non_z) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2886 |
hence "g holomorphic_on (s-{z})" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2887 |
apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2888 |
by (auto simp add:g_def) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2889 |
ultimately show ?thesis unfolding g_def using \<open>open s\<close> |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2890 |
by (auto elim!: no_isolated_singularity) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2891 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2892 |
|
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2893 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2894 |
(*order of the zero of f at z*) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2895 |
definition zorder::"(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> nat" where |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2896 |
"zorder f z = (THE n. n>0 \<and> (\<exists>h r. r>0 \<and> h holomorphic_on cball z r |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2897 |
\<and> (\<forall>w\<in>cball z r. f w = h w * (w-z)^n \<and> h w \<noteq>0)))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2898 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2899 |
definition zer_poly::"[complex \<Rightarrow> complex,complex]\<Rightarrow>complex \<Rightarrow> complex" where |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2900 |
"zer_poly f z = (SOME h. \<exists>r . r>0 \<and> h holomorphic_on cball z r |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2901 |
\<and> (\<forall>w\<in>cball z r. f w = h w * (w-z)^(zorder f z) \<and> h w \<noteq>0))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2902 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2903 |
(*order of the pole of f at z*) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2904 |
definition porder::"(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> nat" where |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2905 |
"porder f z = (let f'=(\<lambda>x. if x=z then 0 else inverse (f x)) in zorder f' z)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2906 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2907 |
definition pol_poly::"[complex \<Rightarrow> complex,complex]\<Rightarrow>complex \<Rightarrow> complex" where |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2908 |
"pol_poly f z = (let f'=(\<lambda> x. if x=z then 0 else inverse (f x)) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2909 |
in inverse o zer_poly f' z)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2910 |
|
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2911 |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2912 |
lemma holomorphic_factor_zero_unique: |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2913 |
fixes f::"complex \<Rightarrow> complex" and z::complex and r::real |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2914 |
assumes "r>0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2915 |
and asm:"\<forall>w\<in>ball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0 \<and> f w = (w - z)^m * h w \<and> h w\<noteq>0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2916 |
and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2917 |
shows "n=m" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2918 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2919 |
have "n>m \<Longrightarrow> False" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2920 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2921 |
assume "n>m" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2922 |
have "(h \<longlongrightarrow> 0) (at z within ball z r)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2923 |
proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) ^ (n - m) * g w"]) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2924 |
have "\<forall>w\<in>ball z r. w\<noteq>z \<longrightarrow> h w = (w-z)^(n-m) * g w" using \<open>n>m\<close> asm |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2925 |
by (auto simp add:field_simps power_diff) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2926 |
then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2927 |
\<Longrightarrow> (x' - z) ^ (n - m) * g x' = h x'" for x' by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2928 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2929 |
define F where "F \<equiv> at z within ball z r" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2930 |
define f' where "f' \<equiv> \<lambda>x. (x - z) ^ (n-m)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2931 |
have "f' z=0" using \<open>n>m\<close> unfolding f'_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2932 |
moreover have "continuous F f'" unfolding f'_def F_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2933 |
by (intro continuous_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2934 |
ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2935 |
by (simp add: continuous_within) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2936 |
moreover have "(g \<longlongrightarrow> g z) F" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2937 |
using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \<open>r>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2938 |
unfolding F_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2939 |
ultimately show " ((\<lambda>w. f' w * g w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2940 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2941 |
moreover have "(h \<longlongrightarrow> h z) (at z within ball z r)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2942 |
using holomorphic_on_imp_continuous_on[OF h_holo] |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2943 |
by (auto simp add:continuous_on_def \<open>r>0\<close>) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2944 |
moreover have "at z within ball z r \<noteq> bot" using \<open>r>0\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2945 |
by (auto simp add:trivial_limit_within islimpt_ball) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2946 |
ultimately have "h z=0" by (auto intro: tendsto_unique) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2947 |
thus False using asm \<open>r>0\<close> by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2948 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2949 |
moreover have "m>n \<Longrightarrow> False" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2950 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2951 |
assume "m>n" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2952 |
have "(g \<longlongrightarrow> 0) (at z within ball z r)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2953 |
proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) ^ (m - n) * h w"]) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2954 |
have "\<forall>w\<in>ball z r. w\<noteq>z \<longrightarrow> g w = (w-z)^(m-n) * h w" using \<open>m>n\<close> asm |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2955 |
by (auto simp add:field_simps power_diff) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2956 |
then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2957 |
\<Longrightarrow> (x' - z) ^ (m - n) * h x' = g x'" for x' by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2958 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2959 |
define F where "F \<equiv> at z within ball z r" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2960 |
define f' where "f' \<equiv>\<lambda>x. (x - z) ^ (m-n)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2961 |
have "f' z=0" using \<open>m>n\<close> unfolding f'_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2962 |
moreover have "continuous F f'" unfolding f'_def F_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2963 |
by (intro continuous_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2964 |
ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2965 |
by (simp add: continuous_within) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2966 |
moreover have "(h \<longlongrightarrow> h z) F" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2967 |
using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \<open>r>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2968 |
unfolding F_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2969 |
ultimately show " ((\<lambda>w. f' w * h w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2970 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2971 |
moreover have "(g \<longlongrightarrow> g z) (at z within ball z r)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2972 |
using holomorphic_on_imp_continuous_on[OF g_holo] |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2973 |
by (auto simp add:continuous_on_def \<open>r>0\<close>) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2974 |
moreover have "at z within ball z r \<noteq> bot" using \<open>r>0\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2975 |
by (auto simp add:trivial_limit_within islimpt_ball) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2976 |
ultimately have "g z=0" by (auto intro: tendsto_unique) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2977 |
thus False using asm \<open>r>0\<close> by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2978 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2979 |
ultimately show "n=m" by fastforce |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2980 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2981 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
2982 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2983 |
lemma holomorphic_factor_zero_Ex1: |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2984 |
assumes "open s" "connected s" "z \<in> s" and |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2985 |
holo:"f holomorphic_on s" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2986 |
and "f z = 0" and "\<exists>w\<in>s. f w \<noteq> 0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2987 |
shows "\<exists>!n. \<exists>g r. 0 < n \<and> 0 < r \<and> |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2988 |
g holomorphic_on cball z r |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2989 |
\<and> (\<forall>w\<in>cball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2990 |
proof (rule ex_ex1I) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2991 |
obtain g r n where "0 < n" "0 < r" "ball z r \<subseteq> s" and |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2992 |
g:"g holomorphic_on ball z r" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2993 |
"\<And>w. w \<in> ball z r \<Longrightarrow> f w = (w - z) ^ n * g w" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2994 |
"\<And>w. w \<in> ball z r \<Longrightarrow> g w \<noteq> 0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2995 |
using holomorphic_factor_zero_nonconstant[OF holo \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> \<open>f z=0\<close>] |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2996 |
by (metis assms(3) assms(5) assms(6)) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2997 |
def r'\<equiv>"r/2" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2998 |
have "cball z r' \<subseteq> ball z r" unfolding r'_def by (simp add: \<open>0 < r\<close> cball_subset_ball_iff) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
2999 |
hence "cball z r' \<subseteq> s" "g holomorphic_on cball z r'" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3000 |
"(\<forall>w\<in>cball z r'. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3001 |
using g \<open>ball z r \<subseteq> s\<close> by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3002 |
moreover have "r'>0" unfolding r'_def using \<open>0<r\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3003 |
ultimately show "\<exists>n g r. 0 < n \<and> 0 < r \<and> g holomorphic_on cball z r |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3004 |
\<and> (\<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3005 |
apply (intro exI[of _ n] exI[of _ g] exI[of _ r']) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3006 |
by (simp add:\<open>0 < n\<close>) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3007 |
next |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3008 |
fix m n |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3009 |
define fac where "fac \<equiv> \<lambda>n g r. \<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3010 |
assume n_asm:"\<exists>g r1. 0 < n \<and> 0 < r1 \<and> g holomorphic_on cball z r1 \<and> fac n g r1" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3011 |
and m_asm:"\<exists>h r2. 0 < m \<and> 0 < r2 \<and> h holomorphic_on cball z r2 \<and> fac m h r2" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3012 |
obtain g r1 where "0 < n" "0 < r1" and g_holo: "g holomorphic_on cball z r1" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3013 |
and "fac n g r1" using n_asm by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3014 |
obtain h r2 where "0 < m" "0 < r2" and h_holo: "h holomorphic_on cball z r2" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3015 |
and "fac m h r2" using m_asm by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3016 |
define r where "r \<equiv> min r1 r2" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3017 |
have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3018 |
moreover have "\<forall>w\<in>ball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0 \<and> f w = (w - z)^m * h w \<and> h w\<noteq>0" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3019 |
using \<open>fac m h r2\<close> \<open>fac n g r1\<close> unfolding fac_def r_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3020 |
by fastforce |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3021 |
ultimately show "m=n" using g_holo h_holo |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3022 |
apply (elim holomorphic_factor_zero_unique[of r z f n g m h,symmetric,rotated]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3023 |
by (auto simp add:r_def) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3024 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3025 |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3026 |
lemma zorder_exist: |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3027 |
fixes f::"complex \<Rightarrow> complex" and z::complex |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3028 |
defines "n\<equiv>zorder f z" and "h\<equiv>zer_poly f z" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3029 |
assumes "open s" "connected s" "z\<in>s" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3030 |
and holo: "f holomorphic_on s" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3031 |
and "f z=0" "\<exists>w\<in>s. f w\<noteq>0" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3032 |
shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3033 |
\<and> (\<forall>w\<in>cball z r. f w = h w * (w-z)^n \<and> h w \<noteq>0) " |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3034 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3035 |
define P where "P \<equiv> \<lambda>h r n. r>0 \<and> h holomorphic_on cball z r |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3036 |
\<and> (\<forall>w\<in>cball z r. ( f w = h w * (w-z)^n) \<and> h w \<noteq>0)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3037 |
have "(\<exists>!n. n>0 \<and> (\<exists> h r. P h r n))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3038 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3039 |
have "\<exists>!n. \<exists>h r. n>0 \<and> P h r n" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3040 |
using holomorphic_factor_zero_Ex1[OF \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> holo \<open>f z=0\<close> |
62837 | 3041 |
\<open>\<exists>w\<in>s. f w\<noteq>0\<close>] unfolding P_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3042 |
apply (subst mult.commute) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3043 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3044 |
thus ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3045 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3046 |
moreover have n:"n=(THE n. n>0 \<and> (\<exists>h r. P h r n))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3047 |
unfolding n_def zorder_def P_def by simp |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3048 |
ultimately have "n>0 \<and> (\<exists>h r. P h r n)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3049 |
apply (drule_tac theI') |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3050 |
by simp |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3051 |
then have "n>0" and "\<exists>h r. P h r n" by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3052 |
moreover have "h=(SOME h. \<exists>r. P h r n)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3053 |
unfolding h_def P_def zer_poly_def[of f z,folded n_def P_def] by simp |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3054 |
ultimately have "\<exists>r. P h r n" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3055 |
apply (drule_tac someI_ex) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3056 |
by simp |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3057 |
then obtain r1 where "P h r1 n" by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3058 |
obtain r2 where "r2>0" "cball z r2 \<subseteq> s" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3059 |
using assms(3) assms(5) open_contains_cball_eq by blast |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3060 |
define r3 where "r3 \<equiv> min r1 r2" |
62837 | 3061 |
have "P h r3 n" using \<open>P h r1 n\<close> \<open>r2>0\<close> unfolding P_def r3_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3062 |
by auto |
62837 | 3063 |
moreover have "cball z r3 \<subseteq> s" using \<open>cball z r2 \<subseteq> s\<close> unfolding r3_def by auto |
3064 |
ultimately show ?thesis using \<open>n>0\<close> unfolding P_def by auto |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3065 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3066 |
|
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3067 |
lemma porder_exist: |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3068 |
fixes f::"complex \<Rightarrow> complex" and z::complex |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3069 |
defines "n \<equiv> porder f z" and "h \<equiv> pol_poly f z" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3070 |
assumes "open s" "z \<in> s" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3071 |
and holo:"f holomorphic_on s-{z}" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3072 |
and "is_pole f z" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3073 |
shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3074 |
\<and> (\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w-z)^n) \<and> h w \<noteq>0)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3075 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3076 |
obtain e where "e>0" and e_ball:"ball z e \<subseteq> s"and e_def: "\<forall>x\<in>ball z e-{z}. f x\<noteq>0" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3077 |
proof - |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3078 |
have "\<forall>\<^sub>F z in at z. f z \<noteq> 0" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3079 |
using \<open>is_pole f z\<close> filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3080 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3081 |
then obtain e1 where "e1>0" and e1_def: "\<forall>x. x \<noteq> z \<and> dist x z < e1 \<longrightarrow> f x \<noteq> 0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3082 |
using eventually_at[of "\<lambda>x. f x\<noteq>0" z,simplified] by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3083 |
obtain e2 where "e2>0" and "ball z e2 \<subseteq>s" using \<open>open s\<close> \<open>z\<in>s\<close> openE by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3084 |
define e where "e \<equiv> min e1 e2" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3085 |
have "e>0" using \<open>e1>0\<close> \<open>e2>0\<close> unfolding e_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3086 |
moreover have "ball z e \<subseteq> s" unfolding e_def using \<open>ball z e2 \<subseteq> s\<close> by auto |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3087 |
moreover have "\<forall>x\<in>ball z e-{z}. f x\<noteq>0" using e1_def \<open>e1>0\<close> \<open>e2>0\<close> unfolding e_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3088 |
by (simp add: DiffD1 DiffD2 dist_commute singletonI) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3089 |
ultimately show ?thesis using that by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3090 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3091 |
define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3092 |
define zo where "zo \<equiv> zorder g z" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3093 |
define zp where "zp \<equiv> zer_poly g z" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3094 |
have "\<exists>w\<in>ball z e. g w \<noteq> 0" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3095 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3096 |
obtain w where w:"w\<in>ball z e-{z}" using \<open>0 < e\<close> |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3097 |
by (metis open_ball all_not_in_conv centre_in_ball insert_Diff_single |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3098 |
insert_absorb not_open_singleton) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3099 |
hence "w\<noteq>z" "f w\<noteq>0" using e_def[rule_format,of w] mem_ball |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3100 |
by (auto simp add:dist_commute) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3101 |
then show ?thesis unfolding g_def using w by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3102 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3103 |
moreover have "g holomorphic_on ball z e" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3104 |
apply (intro is_pole_inverse_holomorphic[of "ball z e",OF _ _ \<open>is_pole f z\<close> e_def,folded g_def]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3105 |
using holo e_ball by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3106 |
moreover have "g z=0" unfolding g_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3107 |
ultimately obtain r where "0 < zo" "0 < r" "cball z r \<subseteq> ball z e" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3108 |
and zp_holo: "zp holomorphic_on cball z r" and |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3109 |
zp_fac: "\<forall>w\<in>cball z r. g w = zp w * (w - z) ^ zo \<and> zp w \<noteq> 0" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3110 |
using zorder_exist[of "ball z e" z g,simplified,folded zo_def zp_def] \<open>e>0\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3111 |
by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3112 |
have n:"n=zo" and h:"h=inverse o zp" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3113 |
unfolding n_def zo_def porder_def h_def zp_def pol_poly_def g_def by simp_all |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3114 |
have "h holomorphic_on cball z r" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3115 |
using zp_holo zp_fac holomorphic_on_inverse unfolding h comp_def by blast |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3116 |
moreover have "\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w-z)^n) \<and> h w \<noteq>0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3117 |
using zp_fac unfolding h n comp_def g_def |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3118 |
by (metis divide_inverse_commute field_class.field_inverse_zero inverse_inverse_eq |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3119 |
inverse_mult_distrib mult.commute) |
62837 | 3120 |
moreover have "0 < n" unfolding n using \<open>zo>0\<close> by simp |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3121 |
ultimately show ?thesis using \<open>0 < r\<close> \<open>cball z r \<subseteq> ball z e\<close> e_ball by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3122 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3123 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3124 |
lemma residue_porder: |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3125 |
fixes f::"complex \<Rightarrow> complex" and z::complex |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3126 |
defines "n \<equiv> porder f z" and "h \<equiv> pol_poly f z" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3127 |
assumes "open s" "z \<in> s" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3128 |
and holo:"f holomorphic_on s - {z}" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3129 |
and pole:"is_pole f z" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3130 |
shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3131 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3132 |
define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3133 |
obtain r where "0 < n" "0 < r" and r_cball:"cball z r \<subseteq> s" and h_holo: "h holomorphic_on cball z r" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3134 |
and h_divide:"(\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w - z) ^ n) \<and> h w \<noteq> 0)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3135 |
using porder_exist[OF \<open>open s\<close> \<open>z \<in> s\<close> holo pole, folded n_def h_def] by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3136 |
have r_nonzero:"\<And>w. w \<in> ball z r - {z} \<Longrightarrow> f w \<noteq> 0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3137 |
using h_divide by simp |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3138 |
define c where "c \<equiv> 2 * pi * \<i>" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3139 |
define der_f where "der_f \<equiv> ((deriv ^^ (n - 1)) h z / fact (n-1))" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3140 |
def h'\<equiv>"\<lambda>u. h u / (u - z) ^ n" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3141 |
have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3142 |
unfolding h'_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3143 |
proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3144 |
folded c_def Suc_pred'[OF \<open>n>0\<close>]]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3145 |
show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3146 |
show "h holomorphic_on ball z r" using h_holo by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3147 |
show " z \<in> ball z r" using \<open>r>0\<close> by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3148 |
qed |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3149 |
then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3150 |
then have "(f has_contour_integral c * der_f) (circlepath z r)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3151 |
proof (elim has_contour_integral_eq) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3152 |
fix x assume "x \<in> path_image (circlepath z r)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3153 |
hence "x\<in>cball z r - {z}" using \<open>r>0\<close> by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3154 |
then show "h' x = f x" using h_divide unfolding h'_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3155 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3156 |
moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3157 |
using base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>r>0\<close> holo r_cball,folded c_def] . |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3158 |
ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3159 |
hence "der_f = residue f z" unfolding c_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3160 |
thus ?thesis unfolding der_f_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3161 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3162 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3163 |
theorem argument_principle: |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3164 |
fixes f::"complex \<Rightarrow> complex" and poles s:: "complex set" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3165 |
defines "zeros\<equiv>{p. f p=0} - poles" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3166 |
assumes "open s" and |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3167 |
"connected s" and |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3168 |
f_holo:"f holomorphic_on s-poles" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3169 |
h_holo:"h holomorphic_on s" and |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3170 |
"valid_path g" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3171 |
loop:"pathfinish g = pathstart g" and |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3172 |
path_img:"path_image g \<subseteq> s - (zeros \<union> poles)" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3173 |
homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3174 |
finite:"finite (zeros \<union> poles)" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3175 |
poles:"\<forall>p\<in>poles. is_pole f p" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3176 |
shows "contour_integral g (\<lambda>x. deriv f x * h x / f x) = 2 * pi * \<i> * |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3177 |
((\<Sum>p\<in>zeros. winding_number g p * h p * zorder f p) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3178 |
- (\<Sum>p\<in>poles. winding_number g p * h p * porder f p))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3179 |
(is "?L=?R") |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3180 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3181 |
define c where "c \<equiv> 2 * complex_of_real pi * \<i> " |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3182 |
define ff where "ff \<equiv> (\<lambda>x. deriv f x * h x / f x)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3183 |
define cont_pole where "cont_pole \<equiv> \<lambda>ff p e. (ff has_contour_integral - c * porder f p * h p) (circlepath p e)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3184 |
define cont_zero where "cont_zero \<equiv> \<lambda>ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3185 |
define avoid where "avoid \<equiv> \<lambda>p e. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> zeros \<union> poles)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3186 |
have "\<exists>e>0. avoid p e \<and> (p\<in>poles \<longrightarrow> cont_pole ff p e) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p e)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3187 |
when "p\<in>s" for p |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3188 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3189 |
obtain e1 where "e1>0" and e1_avoid:"avoid p e1" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3190 |
using finite_cball_avoid[OF \<open>open s\<close> finite] \<open>p\<in>s\<close> unfolding avoid_def by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3191 |
have "\<exists>e2>0. cball p e2 \<subseteq> ball p e1 \<and> cont_pole ff p e2" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3192 |
when "p\<in>poles" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3193 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3194 |
define po where "po \<equiv> porder f p" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3195 |
define pp where "pp \<equiv> pol_poly f p" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3196 |
def f'\<equiv>"\<lambda>w. pp w / (w - p) ^ po" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3197 |
def ff'\<equiv>"(\<lambda>x. deriv f' x * h x / f' x)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3198 |
have "f holomorphic_on ball p e1 - {p}" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3199 |
apply (intro holomorphic_on_subset[OF f_holo]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3200 |
using e1_avoid \<open>p\<in>poles\<close> unfolding avoid_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3201 |
then obtain r where |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3202 |
"0 < po" "r>0" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3203 |
"cball p r \<subseteq> ball p e1" and |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3204 |
pp_holo:"pp holomorphic_on cball p r" and |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3205 |
pp_po:"(\<forall>w\<in>cball p r. (w\<noteq>p \<longrightarrow> f w = pp w / (w - p) ^ po) \<and> pp w \<noteq> 0)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3206 |
using porder_exist[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>] poles \<open>p\<in>poles\<close> |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3207 |
unfolding po_def pp_def |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3208 |
by auto |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3209 |
define e2 where "e2 \<equiv> r/2" |
62837 | 3210 |
have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3211 |
define anal where "anal \<equiv> \<lambda>w. deriv pp w * h w / pp w" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3212 |
define prin where "prin \<equiv> \<lambda>w. - of_nat po * h w / (w - p)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3213 |
have "((\<lambda>w. prin w + anal w) has_contour_integral - c * po * h p) (circlepath p e2)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3214 |
proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3215 |
have "ball p r \<subseteq> s" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3216 |
using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e1_avoid by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3217 |
then have "cball p e2 \<subseteq> s" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3218 |
using \<open>r>0\<close> unfolding e2_def by auto |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3219 |
then have "(\<lambda>w. - of_nat po * h w) holomorphic_on cball p e2" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3220 |
using h_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3221 |
by (auto intro!: holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3222 |
then show "(prin has_contour_integral - c * of_nat po * h p ) (circlepath p e2)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3223 |
using Cauchy_integral_circlepath_simple[folded c_def, of "\<lambda>w. - of_nat po * h w"] |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3224 |
\<open>e2>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3225 |
unfolding prin_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3226 |
by (auto simp add: mult.assoc) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3227 |
have "anal holomorphic_on ball p r" unfolding anal_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3228 |
using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3229 |
by (auto intro!: holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3230 |
then show "(anal has_contour_integral 0) (circlepath p e2)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3231 |
using e2_def \<open>r>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3232 |
by (auto elim!: Cauchy_theorem_disc_simple) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3233 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3234 |
then have "cont_pole ff' p e2" unfolding cont_pole_def po_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3235 |
proof (elim has_contour_integral_eq) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3236 |
fix w assume "w \<in> path_image (circlepath p e2)" |
62837 | 3237 |
then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3238 |
define wp where "wp \<equiv> w-p" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3239 |
have "wp\<noteq>0" and "pp w \<noteq>0" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3240 |
unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3241 |
moreover have der_f':"deriv f' w = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3242 |
proof (rule DERIV_imp_deriv) |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3243 |
define der where "der \<equiv> - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" |
62837 | 3244 |
have po:"po = Suc (po - Suc 0) " using \<open>po>0\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3245 |
have "(pp has_field_derivative (deriv pp w)) (at w)" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3246 |
using DERIV_deriv_iff_has_field_derivative pp_holo \<open>w\<noteq>p\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3247 |
by (meson open_ball \<open>w \<in> ball p r\<close> ball_subset_cball holomorphic_derivI holomorphic_on_subset) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3248 |
then show "(f' has_field_derivative der) (at w)" |
62837 | 3249 |
using \<open>w\<noteq>p\<close> \<open>po>0\<close> unfolding der_def f'_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3250 |
apply (auto intro!: derivative_eq_intros simp add:field_simps) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3251 |
apply (subst (4) po) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3252 |
apply (subst power_Suc) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3253 |
by (auto simp add:field_simps) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3254 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3255 |
ultimately show "prin w + anal w = ff' w" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3256 |
unfolding ff'_def prin_def anal_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3257 |
apply simp |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3258 |
apply (unfold f'_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3259 |
apply (fold wp_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3260 |
by (auto simp add:field_simps) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3261 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3262 |
then have "cont_pole ff p e2" unfolding cont_pole_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3263 |
proof (elim has_contour_integral_eq) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3264 |
fix w assume "w \<in> path_image (circlepath p e2)" |
62837 | 3265 |
then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3266 |
have "deriv f' w = deriv f w" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3267 |
proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3268 |
show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3269 |
by (auto intro!: holomorphic_intros) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3270 |
next |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3271 |
have "ball p e1 - {p} \<subseteq> s - poles" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3272 |
using avoid_def ball_subset_cball e1_avoid |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3273 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3274 |
then have "ball p r - {p} \<subseteq> s - poles" using \<open>cball p r \<subseteq> ball p e1\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3275 |
using ball_subset_cball by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3276 |
then show "f holomorphic_on ball p r - {p}" using f_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3277 |
by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3278 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3279 |
show "open (ball p r - {p})" by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3280 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3281 |
show "w \<in> ball p r - {p}" using \<open>w\<in>ball p r\<close> \<open>w\<noteq>p\<close> by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3282 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3283 |
fix x assume "x \<in> ball p r - {p}" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3284 |
then show "f' x = f x" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3285 |
using pp_po unfolding f'_def by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3286 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3287 |
moreover have " f' w = f w " |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3288 |
using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po \<open>w\<noteq>p\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3289 |
unfolding f'_def by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3290 |
ultimately show "ff' w = ff w" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3291 |
unfolding ff'_def ff_def by simp |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3292 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3293 |
moreover have "cball p e2 \<subseteq> ball p e1" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3294 |
using \<open>0 < r\<close> \<open>cball p r \<subseteq> ball p e1\<close> e2_def by auto |
62837 | 3295 |
ultimately show ?thesis using \<open>e2>0\<close> by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3296 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3297 |
then obtain e2 where e2:"p\<in>poles \<longrightarrow> e2>0 \<and> cball p e2 \<subseteq> ball p e1 \<and> cont_pole ff p e2" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3298 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3299 |
have "\<exists>e3>0. cball p e3 \<subseteq> ball p e1 \<and> cont_zero ff p e3" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3300 |
when "p\<in>zeros" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3301 |
proof - |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3302 |
define zo where "zo \<equiv> zorder f p" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3303 |
define zp where "zp \<equiv> zer_poly f p" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3304 |
def f'\<equiv>"\<lambda>w. zp w * (w - p) ^ zo" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3305 |
def ff'\<equiv>"(\<lambda>x. deriv f' x * h x / f' x)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3306 |
have "f holomorphic_on ball p e1" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3307 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3308 |
have "ball p e1 \<subseteq> s - poles" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3309 |
using avoid_def ball_subset_cball e1_avoid that zeros_def by fastforce |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3310 |
thus ?thesis using f_holo by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3311 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3312 |
moreover have "f p = 0" using \<open>p\<in>zeros\<close> |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3313 |
using DiffD1 mem_Collect_eq zeros_def by blast |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3314 |
moreover have "\<exists>w\<in>ball p e1. f w \<noteq> 0" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3315 |
proof - |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3316 |
def p'\<equiv>"p+e1/2" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3317 |
have "p'\<in>ball p e1" and "p'\<noteq>p" using \<open>e1>0\<close> unfolding p'_def by (auto simp add:dist_norm) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3318 |
then show "\<exists>w\<in>ball p e1. f w \<noteq> 0" using e1_avoid unfolding avoid_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3319 |
apply (rule_tac x=p' in bexI) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3320 |
by (auto simp add:zeros_def) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3321 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3322 |
ultimately obtain r where |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3323 |
"0 < zo" "r>0" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3324 |
"cball p r \<subseteq> ball p e1" and |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3325 |
pp_holo:"zp holomorphic_on cball p r" and |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3326 |
pp_po:"(\<forall>w\<in>cball p r. f w = zp w * (w - p) ^ zo \<and> zp w \<noteq> 0)" |
62837 | 3327 |
using zorder_exist[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>] unfolding zo_def zp_def |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3328 |
by auto |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3329 |
define e2 where "e2 \<equiv> r/2" |
62837 | 3330 |
have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3331 |
define anal where "anal \<equiv> \<lambda>w. deriv zp w * h w / zp w" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3332 |
define prin where "prin \<equiv> \<lambda>w. of_nat zo * h w / (w - p)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3333 |
have "((\<lambda>w. prin w + anal w) has_contour_integral c * zo * h p) (circlepath p e2)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3334 |
proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3335 |
have "ball p r \<subseteq> s" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3336 |
using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e1_avoid by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3337 |
then have "cball p e2 \<subseteq> s" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3338 |
using \<open>r>0\<close> unfolding e2_def by auto |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3339 |
then have "(\<lambda>w. of_nat zo * h w) holomorphic_on cball p e2" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3340 |
using h_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3341 |
by (auto intro!: holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3342 |
then show "(prin has_contour_integral c * of_nat zo * h p ) (circlepath p e2)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3343 |
using Cauchy_integral_circlepath_simple[folded c_def, of "\<lambda>w. of_nat zo * h w"] |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3344 |
\<open>e2>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3345 |
unfolding prin_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3346 |
by (auto simp add: mult.assoc) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3347 |
have "anal holomorphic_on ball p r" unfolding anal_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3348 |
using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3349 |
by (auto intro!: holomorphic_intros) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3350 |
then show "(anal has_contour_integral 0) (circlepath p e2)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3351 |
using e2_def \<open>r>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3352 |
by (auto elim!: Cauchy_theorem_disc_simple) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3353 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3354 |
then have "cont_zero ff' p e2" unfolding cont_zero_def zo_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3355 |
proof (elim has_contour_integral_eq) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3356 |
fix w assume "w \<in> path_image (circlepath p e2)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3357 |
then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3358 |
define wp where "wp \<equiv> w-p" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3359 |
have "wp\<noteq>0" and "zp w \<noteq>0" |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3360 |
unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3361 |
moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3362 |
proof (rule DERIV_imp_deriv) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3363 |
define der where "der \<equiv> zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3364 |
have po:"zo = Suc (zo - Suc 0) " using \<open>zo>0\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3365 |
have "(zp has_field_derivative (deriv zp w)) (at w)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3366 |
using DERIV_deriv_iff_has_field_derivative pp_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3367 |
by (meson Topology_Euclidean_Space.open_ball \<open>w \<in> ball p r\<close> ball_subset_cball holomorphic_derivI holomorphic_on_subset) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3368 |
then show "(f' has_field_derivative der) (at w)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3369 |
using \<open>w\<noteq>p\<close> \<open>zo>0\<close> unfolding der_def f'_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3370 |
by (auto intro!: derivative_eq_intros simp add:field_simps) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3371 |
qed |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3372 |
ultimately show "prin w + anal w = ff' w" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3373 |
unfolding ff'_def prin_def anal_def |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3374 |
apply simp |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3375 |
apply (unfold f'_def) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3376 |
apply (fold wp_def) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3377 |
apply (auto simp add:field_simps) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3378 |
by (metis Suc_diff_Suc minus_nat.diff_0 power_Suc) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3379 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3380 |
then have "cont_zero ff p e2" unfolding cont_zero_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3381 |
proof (elim has_contour_integral_eq) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3382 |
fix w assume "w \<in> path_image (circlepath p e2)" |
62837 | 3383 |
then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3384 |
have "deriv f' w = deriv f w" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3385 |
proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3386 |
show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3387 |
by (auto intro!: holomorphic_intros) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3388 |
next |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3389 |
have "ball p e1 - {p} \<subseteq> s - poles" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3390 |
using avoid_def ball_subset_cball e1_avoid by auto |
62837 | 3391 |
then have "ball p r - {p} \<subseteq> s - poles" using \<open>cball p r \<subseteq> ball p e1\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3392 |
using ball_subset_cball by blast |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3393 |
then show "f holomorphic_on ball p r - {p}" using f_holo |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3394 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3395 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3396 |
show "open (ball p r - {p})" by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3397 |
next |
62837 | 3398 |
show "w \<in> ball p r - {p}" using \<open>w\<in>ball p r\<close> \<open>w\<noteq>p\<close> by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3399 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3400 |
fix x assume "x \<in> ball p r - {p}" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3401 |
then show "f' x = f x" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3402 |
using pp_po unfolding f'_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3403 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3404 |
moreover have " f' w = f w " |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3405 |
using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po unfolding f'_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3406 |
ultimately show "ff' w = ff w" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3407 |
unfolding ff'_def ff_def by simp |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3408 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3409 |
moreover have "cball p e2 \<subseteq> ball p e1" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3410 |
using \<open>0 < r\<close> \<open>cball p r \<subseteq> ball p e1\<close> e2_def by auto |
62837 | 3411 |
ultimately show ?thesis using \<open>e2>0\<close> by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3412 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3413 |
then obtain e3 where e3:"p\<in>zeros \<longrightarrow> e3>0 \<and> cball p e3 \<subseteq> ball p e1 \<and> cont_zero ff p e3" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3414 |
by auto |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3415 |
define e4 where "e4 \<equiv> if p\<in>poles then e2 else if p\<in>zeros then e3 else e1" |
62837 | 3416 |
have "e4>0" using e2 e3 \<open>e1>0\<close> unfolding e4_def by auto |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3417 |
moreover have "avoid p e4" using e2 e3 \<open>e1>0\<close> e1_avoid unfolding e4_def avoid_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3418 |
moreover have "p\<in>poles \<longrightarrow> cont_pole ff p e4" and "p\<in>zeros \<longrightarrow> cont_zero ff p e4" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3419 |
by (auto simp add: e2 e3 e4_def zeros_def) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3420 |
ultimately show ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3421 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3422 |
then obtain get_e where get_e:"\<forall>p\<in>s. get_e p>0 \<and> avoid p (get_e p) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3423 |
\<and> (p\<in>poles \<longrightarrow> cont_pole ff p (get_e p)) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p (get_e p))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3424 |
by metis |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3425 |
define cont where "cont \<equiv> \<lambda>p. contour_integral (circlepath p (get_e p)) ff" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3426 |
define w where "w \<equiv> \<lambda>p. winding_number g p" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3427 |
have "contour_integral g ff = (\<Sum>p\<in>zeros \<union> poles. w p * cont p)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3428 |
unfolding cont_def w_def |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3429 |
proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected s\<close> finite _ \<open>valid_path g\<close> loop |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3430 |
path_img homo]) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3431 |
have "open (s - (zeros \<union> poles))" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3432 |
then show "ff holomorphic_on s - (zeros \<union> poles)" unfolding ff_def using f_holo h_holo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3433 |
by (auto intro!: holomorphic_intros simp add:zeros_def) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3434 |
next |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3435 |
show "\<forall>p\<in>s. 0 < get_e p \<and> (\<forall>w\<in>cball p (get_e p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> zeros \<union> poles))" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3436 |
using get_e using avoid_def by blast |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3437 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3438 |
also have "... = (\<Sum>p\<in>zeros. w p * cont p) + (\<Sum>p\<in>poles. w p * cont p)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3439 |
using finite |
64267 | 3440 |
apply (subst sum.union_disjoint) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3441 |
by (auto simp add:zeros_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3442 |
also have "... = c * ((\<Sum>p\<in>zeros. w p * h p * zorder f p) - (\<Sum>p\<in>poles. w p * h p * porder f p))" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3443 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3444 |
have "(\<Sum>p\<in>zeros. w p * cont p) = (\<Sum>p\<in>zeros. c * w p * h p * zorder f p)" |
64267 | 3445 |
proof (rule sum.cong[of zeros zeros,simplified]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3446 |
fix p assume "p \<in> zeros" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3447 |
show "w p * cont p = c * w p * h p * (zorder f p)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3448 |
proof (cases "p\<in>s") |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3449 |
assume "p \<in> s" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3450 |
have "cont p = c * h p * (zorder f p)" unfolding cont_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3451 |
apply (rule contour_integral_unique) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3452 |
using get_e \<open>p\<in>s\<close> \<open>p\<in>zeros\<close> unfolding cont_zero_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3453 |
by (metis mult.assoc mult.commute) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3454 |
thus ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3455 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3456 |
assume "p\<notin>s" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3457 |
then have "w p=0" using homo unfolding w_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3458 |
then show ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3459 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3460 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3461 |
then have "(\<Sum>p\<in>zeros. w p * cont p) = c * (\<Sum>p\<in>zeros. w p * h p * zorder f p)" |
64267 | 3462 |
apply (subst sum_distrib_left) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3463 |
by (simp add:algebra_simps) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3464 |
moreover have "(\<Sum>p\<in>poles. w p * cont p) = (\<Sum>p\<in>poles. - c * w p * h p * porder f p)" |
64267 | 3465 |
proof (rule sum.cong[of poles poles,simplified]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3466 |
fix p assume "p \<in> poles" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3467 |
show "w p * cont p = - c * w p * h p * (porder f p)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3468 |
proof (cases "p\<in>s") |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3469 |
assume "p \<in> s" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3470 |
have "cont p = - c * h p * (porder f p)" unfolding cont_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3471 |
apply (rule contour_integral_unique) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3472 |
using get_e \<open>p\<in>s\<close> \<open>p\<in>poles\<close> unfolding cont_pole_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3473 |
by (metis mult.assoc mult.commute) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3474 |
thus ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3475 |
next |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3476 |
assume "p\<notin>s" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3477 |
then have "w p=0" using homo unfolding w_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3478 |
then show ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3479 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3480 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3481 |
then have "(\<Sum>p\<in>poles. w p * cont p) = - c * (\<Sum>p\<in>poles. w p * h p * porder f p)" |
64267 | 3482 |
apply (subst sum_distrib_left) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3483 |
by (simp add:algebra_simps) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3484 |
ultimately show ?thesis by (simp add: right_diff_distrib) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3485 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3486 |
finally show ?thesis unfolding w_def ff_def c_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3487 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3488 |
|
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3489 |
subsection \<open>Rouche's theorem \<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3490 |
|
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3491 |
theorem Rouche_theorem: |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3492 |
fixes f g::"complex \<Rightarrow> complex" and s:: "complex set" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3493 |
defines "fg\<equiv>(\<lambda>p. f p+ g p)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3494 |
defines "zeros_fg\<equiv>{p. fg p =0}" and "zeros_f\<equiv>{p. f p=0}" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3495 |
assumes |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3496 |
"open s" and "connected s" and |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3497 |
"finite zeros_fg" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3498 |
"finite zeros_f" and |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3499 |
f_holo:"f holomorphic_on s" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3500 |
g_holo:"g holomorphic_on s" and |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3501 |
"valid_path \<gamma>" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3502 |
loop:"pathfinish \<gamma> = pathstart \<gamma>" and |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3503 |
path_img:"path_image \<gamma> \<subseteq> s " and |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3504 |
path_less:"\<forall>z\<in>path_image \<gamma>. cmod(f z) > cmod(g z)" and |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3505 |
homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number \<gamma> z = 0" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3506 |
shows "(\<Sum>p\<in>zeros_fg. winding_number \<gamma> p * zorder fg p) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3507 |
= (\<Sum>p\<in>zeros_f. winding_number \<gamma> p * zorder f p)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3508 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3509 |
have path_fg:"path_image \<gamma> \<subseteq> s - zeros_fg" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3510 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3511 |
have False when "z\<in>path_image \<gamma>" and "f z + g z=0" for z |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3512 |
proof - |
62837 | 3513 |
have "cmod (f z) > cmod (g z)" using \<open>z\<in>path_image \<gamma>\<close> path_less by auto |
3514 |
moreover have "f z = - g z" using \<open>f z + g z =0\<close> by (simp add: eq_neg_iff_add_eq_0) |
|
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3515 |
then have "cmod (f z) = cmod (g z)" by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3516 |
ultimately show False by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3517 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3518 |
then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3519 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3520 |
have path_f:"path_image \<gamma> \<subseteq> s - zeros_f" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3521 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3522 |
have False when "z\<in>path_image \<gamma>" and "f z =0" for z |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3523 |
proof - |
62837 | 3524 |
have "cmod (g z) < cmod (f z) " using \<open>z\<in>path_image \<gamma>\<close> path_less by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3525 |
then have "cmod (g z) < 0" using \<open>f z=0\<close> by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3526 |
then show False by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3527 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3528 |
then show ?thesis unfolding zeros_f_def using path_img by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3529 |
qed |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3530 |
define w where "w \<equiv> \<lambda>p. winding_number \<gamma> p" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3531 |
define c where "c \<equiv> 2 * complex_of_real pi * \<i>" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3532 |
define h where "h \<equiv> \<lambda>p. g p / f p + 1" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3533 |
obtain spikes |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3534 |
where "finite spikes" and spikes: "\<forall>x\<in>{0..1} - spikes. \<gamma> differentiable at x" |
62837 | 3535 |
using \<open>valid_path \<gamma>\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3536 |
by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3537 |
have h_contour:"((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3538 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3539 |
have outside_img:"0 \<in> outside (path_image (h o \<gamma>))" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3540 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3541 |
have "h p \<in> ball 1 1" when "p\<in>path_image \<gamma>" for p |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3542 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3543 |
have "cmod (g p/f p) <1" using path_less[rule_format,OF that] |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3544 |
apply (cases "cmod (f p) = 0") |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3545 |
by (auto simp add: norm_divide) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3546 |
then show ?thesis unfolding h_def by (auto simp add:dist_complex_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3547 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3548 |
then have "path_image (h o \<gamma>) \<subseteq> ball 1 1" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3549 |
by (simp add: image_subset_iff path_image_compose) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3550 |
moreover have " (0::complex) \<notin> ball 1 1" by (simp add: dist_norm) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3551 |
ultimately show "?thesis" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3552 |
using convex_in_outside[of "ball 1 1" 0] outside_mono by blast |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3553 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3554 |
have valid_h:"valid_path (h \<circ> \<gamma>)" |
62837 | 3555 |
proof (rule valid_path_compose_holomorphic[OF \<open>valid_path \<gamma>\<close> _ _ path_f]) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3556 |
show "h holomorphic_on s - zeros_f" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3557 |
unfolding h_def using f_holo g_holo |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3558 |
by (auto intro!: holomorphic_intros simp add:zeros_f_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3559 |
next |
62837 | 3560 |
show "open (s - zeros_f)" using \<open>finite zeros_f\<close> \<open>open s\<close> finite_imp_closed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3561 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3562 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3563 |
have "((\<lambda>z. 1/z) has_contour_integral 0) (h \<circ> \<gamma>)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3564 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3565 |
have "0 \<notin> path_image (h \<circ> \<gamma>)" using outside_img by (simp add: outside_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3566 |
then have "((\<lambda>z. 1/z) has_contour_integral c * winding_number (h \<circ> \<gamma>) 0) (h \<circ> \<gamma>)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3567 |
using has_contour_integral_winding_number[of "h o \<gamma>" 0,simplified] valid_h |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3568 |
unfolding c_def by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3569 |
moreover have "winding_number (h o \<gamma>) 0 = 0" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3570 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3571 |
have "0 \<in> outside (path_image (h \<circ> \<gamma>))" using outside_img . |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3572 |
moreover have "path (h o \<gamma>)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3573 |
using valid_h by (simp add: valid_path_imp_path) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3574 |
moreover have "pathfinish (h o \<gamma>) = pathstart (h o \<gamma>)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3575 |
by (simp add: loop pathfinish_compose pathstart_compose) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3576 |
ultimately show ?thesis using winding_number_zero_in_outside by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3577 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3578 |
ultimately show ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3579 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3580 |
moreover have "vector_derivative (h \<circ> \<gamma>) (at x) = vector_derivative \<gamma> (at x) * deriv h (\<gamma> x)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3581 |
when "x\<in>{0..1} - spikes" for x |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3582 |
proof (rule vector_derivative_chain_at_general) |
62837 | 3583 |
show "\<gamma> differentiable at x" using that \<open>valid_path \<gamma>\<close> spikes by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3584 |
next |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3585 |
define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3586 |
define t where "t \<equiv> \<gamma> x" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3587 |
have "f t\<noteq>0" unfolding zeros_f_def t_def |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3588 |
by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3589 |
moreover have "t\<in>s" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3590 |
using contra_subsetD path_image_def path_fg t_def that by fastforce |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3591 |
ultimately have "(h has_field_derivative der t) (at t)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3592 |
unfolding h_def der_def using g_holo f_holo \<open>open s\<close> |
64394 | 3593 |
by (auto intro!: holomorphic_derivI derivative_eq_intros) |
3594 |
then show "h field_differentiable at (\<gamma> x)" |
|
3595 |
unfolding t_def field_differentiable_def by blast |
|
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3596 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3597 |
then have " (op / 1 has_contour_integral 0) (h \<circ> \<gamma>) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3598 |
= ((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3599 |
unfolding has_contour_integral |
62837 | 3600 |
apply (intro has_integral_spike_eq[OF negligible_finite, OF \<open>finite spikes\<close>]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3601 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3602 |
ultimately show ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3603 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3604 |
then have "contour_integral \<gamma> (\<lambda>x. deriv h x / h x) = 0" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3605 |
using contour_integral_unique by simp |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3606 |
moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = contour_integral \<gamma> (\<lambda>x. deriv f x / f x) |
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3607 |
+ contour_integral \<gamma> (\<lambda>p. deriv h p / h p)" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3608 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3609 |
have "(\<lambda>p. deriv f p / f p) contour_integrable_on \<gamma>" |
62837 | 3610 |
proof (rule contour_integrable_holomorphic_simple[OF _ _ \<open>valid_path \<gamma>\<close> path_f]) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3611 |
show "open (s - zeros_f)" using finite_imp_closed[OF \<open>finite zeros_f\<close>] \<open>open s\<close> |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3612 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3613 |
then show "(\<lambda>p. deriv f p / f p) holomorphic_on s - zeros_f" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3614 |
using f_holo |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3615 |
by (auto intro!: holomorphic_intros simp add:zeros_f_def) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3616 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3617 |
moreover have "(\<lambda>p. deriv h p / h p) contour_integrable_on \<gamma>" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3618 |
using h_contour |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3619 |
by (simp add: has_contour_integral_integrable) |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3620 |
ultimately have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x + deriv h x / h x) = |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3621 |
contour_integral \<gamma> (\<lambda>p. deriv f p / f p) + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)" |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3622 |
using contour_integral_add[of "(\<lambda>p. deriv f p / f p)" \<gamma> "(\<lambda>p. deriv h p / h p)" ] |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3623 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3624 |
moreover have "deriv fg p / fg p = deriv f p / f p + deriv h p / h p" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3625 |
when "p\<in> path_image \<gamma>" for p |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3626 |
proof - |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3627 |
have "fg p\<noteq>0" and "f p\<noteq>0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3628 |
by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3629 |
have "h p\<noteq>0" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3630 |
proof (rule ccontr) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3631 |
assume "\<not> h p \<noteq> 0" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3632 |
then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3633 |
then have "cmod (g p/f p) = 1" by auto |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3634 |
moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3635 |
apply (cases "cmod (f p) = 0") |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3636 |
by (auto simp add: norm_divide) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3637 |
ultimately show False by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3638 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3639 |
have der_fg:"deriv fg p = deriv f p + deriv g p" unfolding fg_def |
62837 | 3640 |
using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _ \<open>open s\<close>] path_img that |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3641 |
by auto |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3642 |
have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3643 |
proof - |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3644 |
define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3645 |
have "p\<in>s" using path_img that by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3646 |
then have "(h has_field_derivative der p) (at p)" |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3647 |
unfolding h_def der_def using g_holo f_holo \<open>open s\<close> \<open>f p\<noteq>0\<close> |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3648 |
by (auto intro!: derivative_eq_intros holomorphic_derivI) |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3649 |
then show ?thesis unfolding der_def using DERIV_imp_deriv by auto |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3650 |
qed |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3651 |
show ?thesis |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3652 |
apply (simp only:der_fg der_h) |
62837 | 3653 |
apply (auto simp add:field_simps \<open>h p\<noteq>0\<close> \<open>f p\<noteq>0\<close> \<open>fg p\<noteq>0\<close>) |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3654 |
by (auto simp add:field_simps h_def \<open>f p\<noteq>0\<close> fg_def) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3655 |
qed |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3656 |
then have "contour_integral \<gamma> (\<lambda>p. deriv fg p / fg p) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3657 |
= contour_integral \<gamma> (\<lambda>p. deriv f p / f p + deriv h p / h p)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3658 |
by (elim contour_integral_eq) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3659 |
ultimately show ?thesis by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3660 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3661 |
moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = c * (\<Sum>p\<in>zeros_fg. w p * zorder fg p)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3662 |
unfolding c_def zeros_fg_def w_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3663 |
proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3664 |
, of _ "{}" "\<lambda>_. 1",simplified]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3665 |
show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3666 |
show "path_image \<gamma> \<subseteq> s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def . |
62837 | 3667 |
show " finite {p. fg p = 0}" using \<open>finite zeros_fg\<close> unfolding zeros_fg_def . |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3668 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3669 |
moreover have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x) = c * (\<Sum>p\<in>zeros_f. w p * zorder f p)" |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3670 |
unfolding c_def zeros_f_def w_def |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3671 |
proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo |
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3672 |
, of _ "{}" "\<lambda>_. 1",simplified]) |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3673 |
show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3674 |
show "path_image \<gamma> \<subseteq> s - {p. f p = 0}" using path_f unfolding zeros_f_def . |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3675 |
show " finite {p. f p = 0}" using \<open>finite zeros_f\<close> unfolding zeros_f_def . |
63594
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents:
63589
diff
changeset
|
3676 |
qed |
63151
82df5181d699
updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
3677 |
ultimately have " c* (\<Sum>p\<in>zeros_fg. w p * (zorder fg p)) = c* (\<Sum>p\<in>zeros_f. w p * (zorder f p))" |
62540
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3678 |
by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3679 |
then show ?thesis unfolding c_def using w_def by auto |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3680 |
qed |
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents:
62534
diff
changeset
|
3681 |
|
62408
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
3682 |
end |