Theory Conformal_Mappings
section ‹Conformal Mappings and Consequences of Cauchy's Integral Theorem›
text‹By John Harrison et al. Ported from HOL Light by L C Paulson (2016)›
text‹Also Cauchy's residue theorem by Wenda Li (2016)›
theory Conformal_Mappings
imports Cauchy_Integral_Formula
begin
subsection ‹Analytic continuation›
proposition isolated_zeros:
assumes holf: "f holomorphic_on S"
and "open S" "connected S" "ξ ∈ S" "f ξ = 0" "β ∈ S" "f β ≠ 0"
obtains r where "0 < r" and "ball ξ r ⊆ S" and
"⋀z. z ∈ ball ξ r - {ξ} ⟹ f z ≠ 0"
proof -
obtain r where "0 < r" and r: "ball ξ r ⊆ S"
using ‹open S› ‹ξ ∈ S› open_contains_ball_eq by blast
have powf: "((λn. (deriv ^^ n) f ξ / (fact n) * (z - ξ)^n) sums f z)" if "z ∈ ball ξ r" for z
by (intro holomorphic_power_series [OF _ that] holomorphic_on_subset [OF holf r])
obtain m where m: "(deriv ^^ m) f ξ / (fact m) ≠ 0"
using holomorphic_fun_eq_0_on_connected [OF holf ‹open S› ‹connected S› _ ‹ξ ∈ S› ‹β ∈ S›] ‹f β ≠ 0›
by auto
then have "m ≠ 0" using assms(5) funpow_0 by fastforce
obtain s where "0 < s" and s: "⋀z. z ∈ cball ξ s - {ξ} ⟹ f z ≠ 0"
using powser_0_nonzero [OF ‹0 < r› powf ‹f ξ = 0› m]
by (metis ‹m ≠ 0› dist_norm mem_ball norm_minus_commute not_gr_zero)
have "0 < min r s" by (simp add: ‹0 < r› ‹0 < s›)
then show thesis
proof qed (use r s in auto)
qed
proposition analytic_continuation:
assumes holf: "f holomorphic_on S"
and "open S" and "connected S"
and "U ⊆ S" and "ξ ∈ S"
and "ξ islimpt U"
and fU0 [simp]: "⋀z. z ∈ U ⟹ f z = 0"
and "w ∈ S"
shows "f w = 0"
proof -
obtain e where "0 < e" and e: "cball ξ e ⊆ S"
using ‹open S› ‹ξ ∈ S› open_contains_cball_eq by blast
define T where "T = cball ξ e ∩ U"
have contf: "continuous_on (closure T) f"
by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on
holomorphic_on_subset inf.cobounded1)
have fT0 [simp]: "⋀x. x ∈ T ⟹ f x = 0"
by (simp add: T_def)
have "⋀r. ⟦∀e>0. ∃x'∈U. x' ≠ ξ ∧ dist x' ξ < e; 0 < r⟧ ⟹ ∃x'∈cball ξ e ∩ U. x' ≠ ξ ∧ dist x' ξ < r"
by (metis ‹0 < e› IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj)
then have "ξ islimpt T" using ‹ξ islimpt U›
by (auto simp: T_def islimpt_approachable)
then have "ξ ∈ closure T"
by (simp add: closure_def)
then have "f ξ = 0"
by (auto simp: continuous_constant_on_closure [OF contf])
moreover have "⋀r. ⟦0 < r; ⋀z. z ∈ ball ξ r - {ξ} ⟹ f z ≠ 0⟧ ⟹ False"
by (metis open_ball ‹ξ islimpt T› centre_in_ball fT0 insertE insert_Diff islimptE)
ultimately show ?thesis
by (metis ‹open S› ‹connected S› ‹ξ ∈ S› ‹w ∈ S› holf isolated_zeros)
qed
corollary analytic_continuation_open:
assumes "open s" and "open s'" and "s ≠ {}" and "connected s'"
and "s ⊆ s'"
assumes "f holomorphic_on s'" and "g holomorphic_on s'"
and "⋀z. z ∈ s ⟹ f z = g z"
assumes "z ∈ s'"
shows "f z = g z"
proof -
from ‹s ≠ {}› obtain ξ where "ξ ∈ s" by auto
with ‹open s› have ξ: "ξ islimpt s"
by (intro interior_limit_point) (auto simp: interior_open)
have "f z - g z = 0"
by (rule analytic_continuation[of "λz. f z - g z" s' s ξ])
(insert assms ‹ξ ∈ s› ξ, auto intro: holomorphic_intros)
thus ?thesis by simp
qed
corollary analytic_continuation':
assumes "f holomorphic_on S" "open S" "connected S"
and "U ⊆ S" "ξ ∈ S" "ξ islimpt U"
and "f constant_on U"
shows "f constant_on S"
proof -
obtain c where c: "⋀x. x ∈ U ⟹ f x - c = 0"
by (metis ‹f constant_on U› constant_on_def diff_self)
have "(λz. f z - c) holomorphic_on S"
using assms by (intro holomorphic_intros)
with c analytic_continuation assms have "⋀x. x ∈ S ⟹ f x - c = 0"
by blast
then show ?thesis
unfolding constant_on_def by force
qed
lemma holomorphic_compact_finite_zeros:
assumes S: "f holomorphic_on S" "open S" "connected S"
and "compact K" "K ⊆ S"
and "¬ f constant_on S"
shows "finite {z∈K. f z = 0}"
proof (rule ccontr)
assume "infinite {z∈K. f z = 0}"
then obtain z where "z ∈ K" and z: "z islimpt {z∈K. f z = 0}"
using ‹compact K› by (auto simp: compact_eq_Bolzano_Weierstrass)
moreover have "{z∈K. f z = 0} ⊆ S"
using ‹K ⊆ S› by blast
ultimately show False
using assms analytic_continuation [OF S] unfolding constant_on_def
by blast
qed
lemma holomorphic_countable_zeros:
assumes S: "f holomorphic_on S" "open S" "connected S" and "fsigma S"
and "¬ f constant_on S"
shows "countable {z∈S. f z = 0}"
proof -
obtain F::"nat ⇒ complex set"
where F: "range F ⊆ Collect compact" and Seq: "S = (⋃i. F i)"
using ‹fsigma S› by (meson fsigma_Union_compact)
have fin: "finite {z ∈ F i. f z = 0}" for i
using holomorphic_compact_finite_zeros assms F Seq Union_iff by blast
have "{z ∈ S. f z = 0} = (⋃i. {z ∈ F i. f z = 0})"
using Seq by auto
with fin show ?thesis
by (simp add: countable_finite)
qed
lemma holomorphic_countable_equal:
assumes "f holomorphic_on S" "g holomorphic_on S" "open S" "connected S" and "fsigma S"
and eq: "uncountable {z∈S. f z = g z}"
shows "S ⊆ {z∈S. f z = g z}"
proof -
obtain z where z: "z∈S" "f z = g z"
using eq not_finite_existsD uncountable_infinite by blast
have "(λx. f x - g x) holomorphic_on S"
by (simp add: assms holomorphic_on_diff)
then have "(λx. f x - g x) constant_on S"
using holomorphic_countable_zeros assms by force
with z have "⋀x. x∈S ⟹ f x - g x = 0"
unfolding constant_on_def by force
then show ?thesis
by auto
qed
lemma holomorphic_countable_equal_UNIV:
assumes fg: "f holomorphic_on UNIV" "g holomorphic_on UNIV"
and eq: "uncountable {z. f z = g z}"
shows "f=g"
using holomorphic_countable_equal [OF fg] eq by fastforce
subsection‹Open mapping theorem›
lemma holomorphic_contract_to_zero:
assumes contf: "continuous_on (cball ξ r) f"
and holf: "f holomorphic_on ball ξ r"
and "0 < r"
and norm_less: "⋀z. norm(ξ - z) = r ⟹ norm(f ξ) < norm(f z)"
obtains z where "z ∈ ball ξ r" "f z = 0"
proof -
{ assume fnz: "⋀w. w ∈ ball ξ r ⟹ f w ≠ 0"
then have "0 < norm (f ξ)"
by (simp add: ‹0 < r›)
have fnz': "⋀w. w ∈ cball ξ r ⟹ f w ≠ 0"
using dist_complex_def fnz norm_less order_le_less by fastforce
have "frontier(cball ξ r) ≠ {}"
using ‹0 < r› by simp
define g where [abs_def]: "g z = inverse (f z)" for z
have contg: "continuous_on (cball ξ r) g"
unfolding g_def using contf continuous_on_inverse fnz' by blast
have holg: "g holomorphic_on ball ξ r"
unfolding g_def using fnz holf holomorphic_on_inverse by blast
have "frontier (cball ξ r) ⊆ cball ξ r"
by (simp add: subset_iff)
then have contf': "continuous_on (frontier (cball ξ r)) f"
and contg': "continuous_on (frontier (cball ξ r)) g"
by (blast intro: contf contg continuous_on_subset)+
have froc: "frontier(cball ξ r) ≠ {}"
using ‹0 < r› by simp
moreover have "continuous_on (frontier (cball ξ r)) (norm o f)"
using contf' continuous_on_compose continuous_on_norm_id by blast
ultimately obtain w where w: "w ∈ frontier(cball ξ r)"
and now: "⋀x. x ∈ frontier(cball ξ r) ⟹ norm (f w) ≤ norm (f x)"
using continuous_attains_inf [OF compact_frontier [OF compact_cball]]
by (metis comp_apply)
then have fw: "0 < norm (f w)"
by (simp add: fnz')
have "continuous_on (frontier (cball ξ r)) (norm o g)"
using contg' continuous_on_compose continuous_on_norm_id by blast
then obtain v where v: "v ∈ frontier(cball ξ r)"
and nov: "⋀x. x ∈ frontier(cball ξ r) ⟹ norm (g v) ≥ norm (g x)"
using continuous_attains_sup [OF compact_frontier [OF compact_cball] froc] by force
then have fv: "0 < norm (f v)"
by (simp add: fnz')
have "norm ((deriv ^^ 0) g ξ) ≤ fact 0 * norm (g v) / r ^ 0"
by (rule Cauchy_inequality [OF holg contg ‹0 < r›]) (simp add: dist_norm nov)
then have "cmod (g ξ) ≤ cmod (g v)"
by simp
moreover have "cmod (ξ - w) = r"
by (metis (no_types) dist_norm frontier_cball mem_sphere w)
ultimately obtain wr: "norm (ξ - w) = r" and nfw: "norm (f w) ≤ norm (f ξ)"
unfolding g_def
by (smt (verit, del_insts) ‹0 < cmod (f ξ)› inverse_le_imp_le norm_inverse now v)
with fw have False
using norm_less by force
}
with that show ?thesis by blast
qed
theorem open_mapping_thm:
assumes holf: "f holomorphic_on S"
and S: "open S" and "connected S"
and "open U" and "U ⊆ S"
and fne: "¬ f constant_on S"
shows "open (f ` U)"
proof -
have *: "open (f ` U)"
if "U ≠ {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "⋀x. ∃y ∈ U. f y ≠ x"
for U
proof (clarsimp simp: open_contains_ball)
fix ξ assume ξ: "ξ ∈ U"
show "∃e>0. ball (f ξ) e ⊆ f ` U"
proof -
have hol: "(λz. f z - f ξ) holomorphic_on U"
by (rule holomorphic_intros that)+
obtain s where "0 < s" and sbU: "ball ξ s ⊆ U"
and sne: "⋀z. z ∈ ball ξ s - {ξ} ⟹ (λz. f z - f ξ) z ≠ 0"
using isolated_zeros [OF hol U ξ] by (metis fneU right_minus_eq)
obtain r where "0 < r" and r: "cball ξ r ⊆ ball ξ s"
using ‹0 < s› by (rule_tac r="s/2" in that) auto
have "cball ξ r ⊆ U"
using sbU r by blast
then have frsbU: "frontier (cball ξ r) ⊆ U"
using Diff_subset frontier_def order_trans by fastforce
then have cof: "compact (frontier(cball ξ r))"
by blast
have frne: "frontier (cball ξ r) ≠ {}"
using ‹0 < r› by auto
have contfr: "continuous_on (frontier (cball ξ r)) (λz. norm (f z - f ξ))"
by (metis continuous_on_norm continuous_on_subset frsbU hol holomorphic_on_imp_continuous_on)
obtain w where "norm (ξ - w) = r"
and w: "(⋀z. norm (ξ - z) = r ⟹ norm (f w - f ξ) ≤ norm(f z - f ξ))"
using continuous_attains_inf [OF cof frne contfr] by (auto simp: dist_norm)
moreover define ε where "ε ≡ norm (f w - f ξ) / 3"
ultimately have "0 < ε"
using ‹0 < r› dist_complex_def r sne by auto
have "ball (f ξ) ε ⊆ f ` U"
proof
fix γ
assume γ: "γ ∈ ball (f ξ) ε"
have *: "cmod (γ - f ξ) < cmod (γ - f z)" if "cmod (ξ - z) = r" for z
proof -
have lt: "cmod (f w - f ξ) / 3 < cmod (γ - f z)"
using w [OF that] γ
using dist_triangle2 [of "f ξ" "γ" "f z"] dist_triangle2 [of "f ξ" "f z" γ]
by (simp add: ε_def dist_norm norm_minus_commute)
show ?thesis
by (metis ε_def dist_commute dist_norm less_trans lt mem_ball γ)
qed
have "continuous_on (cball ξ r) (λz. γ - f z)"
using ‹cball ξ r ⊆ U› ‹f holomorphic_on U›
by (force intro: continuous_intros continuous_on_subset holomorphic_on_imp_continuous_on)
moreover have "(λz. γ - f z) holomorphic_on ball ξ r"
using ‹cball ξ r ⊆ U› ball_subset_cball holomorphic_on_subset that(4)
by (intro holomorphic_intros) blast
ultimately obtain z where "z ∈ ball ξ r" "γ - f z = 0"
using "*" ‹0 < r› holomorphic_contract_to_zero by blast
then show "γ ∈ f ` U"
using ‹cball ξ r ⊆ U› by fastforce
qed
then show ?thesis using ‹0 < ε› by blast
qed
qed
have "open (f ` X)" if "X ∈ components U" for X
proof -
have holfU: "f holomorphic_on U"
using ‹U ⊆ S› holf holomorphic_on_subset by blast
have "X ≠ {}"
using that by (simp add: in_components_nonempty)
moreover have "open X"
using that ‹open U› open_components by auto
moreover have "connected X"
using that in_components_maximal by blast
moreover have "f holomorphic_on X"
by (meson that holfU holomorphic_on_subset in_components_maximal)
moreover have "∃y∈X. f y ≠ x" for x
proof (rule ccontr)
assume not: "¬ (∃y∈X. f y ≠ x)"
have "X ⊆ S"
using ‹U ⊆ S› in_components_subset that by blast
obtain w where w: "w ∈ X" using ‹X ≠ {}› by blast
have wis: "w islimpt X"
using w ‹open X› interior_eq by auto
have hol: "(λz. f z - x) holomorphic_on S"
by (simp add: holf holomorphic_on_diff)
with fne [unfolded constant_on_def]
analytic_continuation[OF hol S ‹connected S› ‹X ⊆ S› _ wis] not ‹X ⊆ S› w
show False by auto
qed
ultimately show ?thesis
by (rule *)
qed
then have "open (f ` ⋃(components U))"
by (metis (no_types, lifting) imageE image_Union open_Union)
then show ?thesis
by force
qed
text‹No need for \<^term>‹S› to be connected. But the nonconstant condition is stronger.›
corollary open_mapping_thm2:
assumes holf: "f holomorphic_on S"
and S: "open S"
and "open U" "U ⊆ S"
and fnc: "⋀X. ⟦open X; X ⊆ S; X ≠ {}⟧ ⟹ ¬ f constant_on X"
shows "open (f ` U)"
proof -
have "S = ⋃(components S)" by simp
with ‹U ⊆ S› have "U = (⋃C ∈ components S. C ∩ U)" by auto
then have "f ` U = (⋃C ∈ components S. f ` (C ∩ U))"
using image_UN by fastforce
moreover
{ fix C assume "C ∈ components S"
with S ‹C ∈ components S› open_components in_components_connected
have C: "open C" "connected C" by auto
have "C ⊆ S"
by (metis ‹C ∈ components S› in_components_maximal)
have nf: "¬ f constant_on C"
using ‹open C› ‹C ∈ components S› ‹C ⊆ S› fnc in_components_nonempty by blast
have "f holomorphic_on C"
by (metis holf holomorphic_on_subset ‹C ⊆ S›)
then have "open (f ` (C ∩ U))"
by (meson C ‹open U› inf_le1 nf open_Int open_mapping_thm)
} ultimately show ?thesis
by force
qed
corollary open_mapping_thm3:
assumes "f holomorphic_on S"
and "open S" and "inj_on f S"
shows "open (f ` S)"
by (meson assms inj_on_subset injective_not_constant open_mapping_thm2 order.refl)
subsection‹Maximum modulus principle›
text‹If \<^term>‹f› is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is
properly within the domain of \<^term>‹f›.›
proposition maximum_modulus_principle:
assumes holf: "f holomorphic_on S"
and S: "open S" and "connected S"
and "open U" and "U ⊆ S" and "ξ ∈ U"
and no: "⋀z. z ∈ U ⟹ norm(f z) ≤ norm(f ξ)"
shows "f constant_on S"
proof (rule ccontr)
assume "¬ f constant_on S"
then have "open (f ` U)"
using open_mapping_thm assms by blast
moreover have "¬ open (f ` U)"
proof -
have "∃t. cmod (f ξ - t) < e ∧ t ∉ f ` U" if "0 < e" for e
using that
apply (rule_tac x="if 0 < Re(f ξ) then f ξ + (e/2) else f ξ - (e/2)" in exI)
apply (simp add: dist_norm)
apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym)
done
then show ?thesis
unfolding open_contains_ball by (metis ‹ξ ∈ U› contra_subsetD dist_norm imageI mem_ball)
qed
ultimately show False
by blast
qed
proposition maximum_modulus_frontier:
assumes holf: "f holomorphic_on (interior S)"
and contf: "continuous_on (closure S) f"
and bos: "bounded S"
and leB: "⋀z. z ∈ frontier S ⟹ norm(f z) ≤ B"
and "ξ ∈ S"
shows "norm(f ξ) ≤ B"
proof -
have "compact (closure S)" using bos
by (simp add: bounded_closure compact_eq_bounded_closed)
moreover have "continuous_on (closure S) (cmod ∘ f)"
using contf continuous_on_compose continuous_on_norm_id by blast
ultimately obtain z where "z ∈ closure S" and z: "⋀y. y ∈ closure S ⟹ (cmod ∘ f) y ≤ (cmod ∘ f) z"
using continuous_attains_sup [of "closure S" "norm o f"] ‹ξ ∈ S› by auto
then consider "z ∈ frontier S" | "z ∈ interior S" using frontier_def by auto
then have "norm(f z) ≤ B"
proof cases
case 1 then show ?thesis using leB by blast
next
case 2
have "f constant_on (connected_component_set (interior S) z)"
proof (rule maximum_modulus_principle)
show "f holomorphic_on connected_component_set (interior S) z"
by (metis connected_component_subset holf holomorphic_on_subset)
show zin: "z ∈ connected_component_set (interior S) z"
by (simp add: 2)
show "⋀W. W ∈ connected_component_set (interior S) z ⟹ cmod (f W) ≤ cmod (f z)"
using closure_def connected_component_subset z by fastforce
qed (auto simp: open_connected_component)
then obtain c where c: "⋀w. w ∈ connected_component_set (interior S) z ⟹ f w = c"
by (auto simp: constant_on_def)
have "f ` closure(connected_component_set (interior S) z) ⊆ {c}"
proof (rule image_closure_subset)
show "continuous_on (closure (connected_component_set (interior S) z)) f"
by (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset)
qed (use c in auto)
then have cc: "⋀w. w ∈ closure(connected_component_set (interior S) z) ⟹ f w = c" by blast
have "connected_component (interior S) z z"
by (simp add: "2")
moreover have "connected_component_set (interior S) z ≠ UNIV"
by (metis bos bounded_interior connected_component_eq_UNIV not_bounded_UNIV)
ultimately have "frontier(connected_component_set (interior S) z) ≠ {}"
by (meson "2" connected_component_eq_empty frontier_not_empty)
then obtain w where w: "w ∈ frontier(connected_component_set (interior S) z)"
by auto
then have "norm (f z) = norm (f w)" by (simp add: "2" c cc frontier_def)
also have "… ≤ B"
using w frontier_interior_subset frontier_of_connected_component_subset
by (blast intro: leB)
finally show ?thesis .
qed
then show ?thesis
using z ‹ξ ∈ S› closure_subset by fastforce
qed
corollary maximum_real_frontier:
assumes holf: "f holomorphic_on (interior S)"
and contf: "continuous_on (closure S) f"
and bos: "bounded S"
and leB: "⋀z. z ∈ frontier S ⟹ Re(f z) ≤ B"
and "ξ ∈ S"
shows "Re(f ξ) ≤ B"
using maximum_modulus_frontier [of "exp o f" S "exp B"]
Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms
by auto
subsection ‹Factoring out a zero according to its order›
lemma holomorphic_factor_order_of_zero:
assumes holf: "f holomorphic_on S"
and os: "open S"
and "ξ ∈ S" "0 < n"
and dnz: "(deriv ^^ n) f ξ ≠ 0"
and dfz: "⋀i. ⟦0 < i; i < n⟧ ⟹ (deriv ^^ i) f ξ = 0"
obtains g r where "0 < r"
"g holomorphic_on ball ξ r"
"⋀w. w ∈ ball ξ r ⟹ f w - f ξ = (w - ξ)^n * g w"
"⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
proof -
obtain r where "r>0" and r: "ball ξ r ⊆ S" using assms by (blast elim!: openE)
then have holfb: "f holomorphic_on ball ξ r"
using holf holomorphic_on_subset by blast
define g where "g w = suminf (λi. (deriv ^^ (i + n)) f ξ / (fact(i + n)) * (w - ξ)^i)" for w
have sumsg: "(λi. (deriv ^^ (i + n)) f ξ / (fact(i + n)) * (w - ξ)^i) sums g w"
and feq: "f w - f ξ = (w - ξ)^n * g w"
if w: "w ∈ ball ξ r" for w
proof -
define powf where "powf = (λi. (deriv ^^ i) f ξ/(fact i) * (w - ξ)^i)"
have [simp]: "powf 0 = f ξ"
by (simp add: powf_def)
have sing: "{..<n} - {i. powf i = 0} = (if f ξ = 0 then {} else {0})"
unfolding powf_def using ‹0 < n› dfz by (auto simp: dfz; metis funpow_0 not_gr0)
have "powf sums f w"
unfolding powf_def by (rule holomorphic_power_series [OF holfb w])
moreover have "(∑i<n. powf i) = f ξ"
by (subst sum.setdiff_irrelevant [symmetric]; simp add: dfz sing)
ultimately have fsums: "(λi. powf (i+n)) sums (f w - f ξ)"
using w sums_iff_shift' by metis
then have *: "summable (λi. (w - ξ) ^ n * ((deriv ^^ (i + n)) f ξ * (w - ξ) ^ i / fact (i + n)))"
unfolding powf_def using sums_summable
by (auto simp: power_add mult_ac)
have "summable (λi. (deriv ^^ (i + n)) f ξ * (w - ξ) ^ i / fact (i + n))"
proof (cases "w=ξ")
case False then show ?thesis
using summable_mult [OF *, of "1 / (w - ξ) ^ n"] by simp
next
case True then show ?thesis
by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"]
split: if_split_asm)
qed
then show sumsg: "(λi. (deriv ^^ (i + n)) f ξ / (fact(i + n)) * (w - ξ)^i) sums g w"
by (simp add: summable_sums_iff g_def)
show "f w - f ξ = (w - ξ)^n * g w"
using sums_mult [OF sumsg, of "(w - ξ) ^ n"]
by (intro sums_unique2 [OF fsums]) (auto simp: power_add mult_ac powf_def)
qed
then have holg: "g holomorphic_on ball ξ r"
by (meson sumsg power_series_holomorphic)
then have contg: "continuous_on (ball ξ r) g"
by (blast intro: holomorphic_on_imp_continuous_on)
have "g ξ ≠ 0"
using dnz unfolding g_def
by (subst suminf_finite [of "{0}"]) auto
obtain d where "0 < d" and d: "⋀w. w ∈ ball ξ d ⟹ g w ≠ 0"
using ‹0 < r› continuous_on_avoid [OF contg _ ‹g ξ ≠ 0›]
by (metis centre_in_ball le_cases mem_ball mem_ball_leI)
show ?thesis
proof
show "g holomorphic_on ball ξ (min r d)"
using holg by (auto simp: feq holomorphic_on_subset subset_ball d)
qed (use ‹0 < r› ‹0 < d› in ‹auto simp: feq d›)
qed
lemma holomorphic_factor_order_of_zero_strong:
assumes holf: "f holomorphic_on S" "open S" "ξ ∈ S" "0 < n"
and "(deriv ^^ n) f ξ ≠ 0"
and "⋀i. ⟦0 < i; i < n⟧ ⟹ (deriv ^^ i) f ξ = 0"
obtains g r where "0 < r"
"g holomorphic_on ball ξ r"
"⋀w. w ∈ ball ξ r ⟹ f w - f ξ = ((w - ξ) * g w) ^ n"
"⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
proof -
obtain g r where "0 < r"
and holg: "g holomorphic_on ball ξ r"
and feq: "⋀w. w ∈ ball ξ r ⟹ f w - f ξ = (w - ξ)^n * g w"
and gne: "⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
by (auto intro: holomorphic_factor_order_of_zero [OF assms])
have con: "continuous_on (ball ξ r) (λz. deriv g z / g z)"
by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on)
have cd: "(λz. deriv g z / g z) field_differentiable at x" if "dist ξ x < r" for x
proof (intro derivative_intros)
show "deriv g field_differentiable at x"
using that holg mem_ball by (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at)
show "g field_differentiable at x"
by (metis that open_ball at_within_open holg holomorphic_on_def mem_ball)
qed (simp add: gne that)
obtain h where h: "⋀x. x ∈ ball ξ r ⟹ (h has_field_derivative deriv g x / g x) (at x)"
using holomorphic_convex_primitive [of "ball ξ r" "{}" "λz. deriv g z / g z"]
by (metis (no_types, lifting) Diff_empty at_within_interior cd con convex_ball infinite_imp_nonempty interior_ball mem_ball)
then have "continuous_on (ball ξ r) h"
by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open)
then have con: "continuous_on (ball ξ r) (λx. exp (h x) / g x)"
by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne)
have gfd: "dist ξ x < r ⟹ g field_differentiable at x" if "dist ξ x < r" for x
using holg holomorphic_on_imp_differentiable_at by auto
have 0: "dist ξ x < r ⟹ ((λx. exp (h x) / g x) has_field_derivative 0) (at x)" for x
by (rule gfd h derivative_eq_intros DERIV_deriv_iff_field_differentiable [THEN iffD2] | simp add: gne)+
obtain c where c: "⋀x. x ∈ ball ξ r ⟹ exp (h x) / g x = c"
by (rule DERIV_zero_connected_constant [of "ball ξ r" "{}" "λx. exp(h x) / g x"]) (auto simp: con 0)
have hol: "(λz. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball ξ r"
proof (intro holomorphic_intros holomorphic_on_compose [unfolded o_def, where g = exp])
show "h holomorphic_on ball ξ r"
using h holomorphic_on_open by blast
qed (use ‹0 < n› in auto)
show ?thesis
proof
show "⋀w. w ∈ ball ξ r ⟹ f w - f ξ = ((w - ξ) * exp ((Ln (inverse c) + h w) / of_nat n)) ^ n"
using ‹0 < n›
by (auto simp: feq power_mult_distrib exp_divide_power_eq exp_add gne simp flip: c)
qed (use hol ‹0 < r› in auto)
qed
lemma
fixes k :: "'a::wellorder"
assumes a_def: "a == LEAST x. P x" and P: "P k"
shows def_LeastI: "P a" and def_Least_le: "a ≤ k"
unfolding a_def
by (rule LeastI Least_le; rule P)+
lemma holomorphic_factor_zero_nonconstant:
assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
and "ξ ∈ S" "f ξ = 0"
and nonconst: "¬ f constant_on S"
obtains g r n
where "0 < n" "0 < r" "ball ξ r ⊆ S"
"g holomorphic_on ball ξ r"
"⋀w. w ∈ ball ξ r ⟹ f w = (w - ξ)^n * g w"
"⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
proof (cases "∀n>0. (deriv ^^ n) f ξ = 0")
case True then show ?thesis
using holomorphic_fun_eq_const_on_connected [OF holf S _ ‹ξ ∈ S›] nonconst by (simp add: constant_on_def)
next
case False
then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f ξ ≠ 0" by blast
obtain r0 where "r0 > 0" "ball ξ r0 ⊆ S" using S openE ‹ξ ∈ S› by auto
define n where "n ≡ LEAST n. (deriv ^^ n) f ξ ≠ 0"
have n_ne: "(deriv ^^ n) f ξ ≠ 0"
by (rule def_LeastI [OF n_def]) (rule n0)
then have "0 < n" using ‹f ξ = 0›
using funpow_0 by fastforce
have n_min: "⋀k. k < n ⟹ (deriv ^^ k) f ξ = 0"
using def_Least_le [OF n_def] not_le by blast
then obtain g r1
where g: "0 < r1" "g holomorphic_on ball ξ r1"
and geq: "⋀w. w ∈ ball ξ r1 ⟹ f w = (w - ξ) ^ n * g w"
and g0: "⋀w. w ∈ ball ξ r1 ⟹ g w ≠ 0"
by (auto intro: holomorphic_factor_order_of_zero [OF holf ‹open S› ‹ξ ∈ S› ‹n > 0› n_ne] simp: ‹f ξ = 0›)
show ?thesis
proof
show "g holomorphic_on ball ξ (min r0 r1)"
using g by auto
show "⋀w. w ∈ ball ξ (min r0 r1) ⟹ f w = (w - ξ) ^ n * g w"
by (simp add: geq)
qed (use ‹0 < n› ‹0 < r0› ‹0 < r1› ‹ball ξ r0 ⊆ S› g0 in auto)
qed
lemma holomorphic_lower_bound_difference:
assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
and "ξ ∈ S" and "φ ∈ S"
and fne: "f φ ≠ f ξ"
obtains k n r
where "0 < k" "0 < r"
"ball ξ r ⊆ S"
"⋀w. w ∈ ball ξ r ⟹ k * norm(w - ξ)^n ≤ norm(f w - f ξ)"
proof -
define n where "n = (LEAST n. 0 < n ∧ (deriv ^^ n) f ξ ≠ 0)"
obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f ξ ≠ 0"
using fne holomorphic_fun_eq_const_on_connected [OF holf S] ‹ξ ∈ S› ‹φ ∈ S› by blast
then have "0 < n" and n_ne: "(deriv ^^ n) f ξ ≠ 0"
unfolding n_def by (metis (mono_tags, lifting) LeastI)+
have n_min: "⋀k. ⟦0 < k; k < n⟧ ⟹ (deriv ^^ k) f ξ = 0"
unfolding n_def by (blast dest: not_less_Least)
then obtain g r
where "0 < r" and holg: "g holomorphic_on ball ξ r"
and fne: "⋀w. w ∈ ball ξ r ⟹ f w - f ξ = (w - ξ) ^ n * g w"
and gnz: "⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
by (auto intro: holomorphic_factor_order_of_zero [OF holf ‹open S› ‹ξ ∈ S› ‹n > 0› n_ne])
obtain e where "e>0" and e: "ball ξ e ⊆ S" using assms by (blast elim!: openE)
then have holfb: "f holomorphic_on ball ξ e"
using holf holomorphic_on_subset by blast
define d where "d = (min e r) / 2"
have "0 < d" using ‹0 < r› ‹0 < e› by (simp add: d_def)
have "d < r"
using ‹0 < r› by (auto simp: d_def)
then have cbb: "cball ξ d ⊆ ball ξ r"
by (auto simp: cball_subset_ball_iff)
then have "g holomorphic_on cball ξ d"
by (rule holomorphic_on_subset [OF holg])
then have "closed (g ` cball ξ d)"
by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on)
moreover have "g ` cball ξ d ≠ {}"
using ‹0 < d› by auto
ultimately obtain x where x: "x ∈ g ` cball ξ d" and "⋀y. y ∈ g ` cball ξ d ⟹ dist 0 x ≤ dist 0 y"
by (rule distance_attains_inf) blast
then have leg: "⋀w. w ∈ cball ξ d ⟹ norm x ≤ norm (g w)"
by auto
have "ball ξ d ⊆ cball ξ d" by auto
also have "… ⊆ ball ξ e" using ‹0 < d› d_def by auto
also have "… ⊆ S" by (rule e)
finally have dS: "ball ξ d ⊆ S" .
have "x ≠ 0" using gnz x ‹d < r› by auto
show thesis
proof
show "⋀w. w ∈ ball ξ d ⟹ cmod x * cmod (w - ξ) ^ n ≤ cmod (f w - f ξ)"
using ‹d < r› leg by (auto simp: fne norm_mult norm_power algebra_simps mult_right_mono)
qed (use dS ‹x ≠ 0› ‹d > 0› in auto)
qed
lemma
assumes holf: "f holomorphic_on (S - {ξ})" and ξ: "ξ ∈ interior S"
shows holomorphic_on_extend_lim:
"(∃g. g holomorphic_on S ∧ (∀z ∈ S - {ξ}. g z = f z)) ⟷
((λz. (z - ξ) * f z) ⤏ 0) (at ξ)"
(is "?P = ?Q")
and holomorphic_on_extend_bounded:
"(∃g. g holomorphic_on S ∧ (∀z ∈ S - {ξ}. g z = f z)) ⟷
(∃B. eventually (λz. norm(f z) ≤ B) (at ξ))"
(is "?P = ?R")
proof -
obtain δ where "0 < δ" and δ: "ball ξ δ ⊆ S"
using ξ mem_interior by blast
have "?R" if holg: "g holomorphic_on S" and gf: "⋀z. z ∈ S - {ξ} ⟹ g z = f z" for g
proof -
have §: "cmod (f x) ≤ cmod (g ξ) + 1" if "x ≠ ξ" "dist x ξ < δ" "dist (g x) (g ξ) < 1" for x
proof -
have "x ∈ S"
by (metis δ dist_commute mem_ball subsetD that(2))
with that gf [of x] show ?thesis
using norm_triangle_ineq2 [of "f x" "g ξ"] dist_complex_def by auto
qed
then have *: "∀⇩F z in at ξ. dist (g z) (g ξ) < 1 ⟶ cmod (f z) ≤ cmod (g ξ) + 1"
using ‹0 < δ› eventually_at by blast
have "continuous_on (interior S) g"
by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset)
then have "⋀x. x ∈ interior S ⟹ (g ⤏ g x) (at x)"
using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast
then have "(g ⤏ g ξ) (at ξ)"
by (simp add: ξ)
then have "∀⇩F z in at ξ. cmod (f z) ≤ cmod (g ξ) + 1"
by (rule eventually_mp [OF * tendstoD [where e=1]], auto)
then show ?thesis
by blast
qed
moreover have "?Q" if "∀⇩F z in at ξ. cmod (f z) ≤ B" for B
by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero)
moreover have "?P" if "(λz. (z - ξ) * f z) ─ξ→ 0"
proof -
define h where [abs_def]: "h z = (z - ξ)^2 * f z" for z
have "(λy. (y - ξ)⇧2 * f y / (y - ξ)) ─ξ→ 0"
by (simp add: LIM_cong power2_eq_square that)
then have h0: "(h has_field_derivative 0) (at ξ)"
by (simp add: h_def has_field_derivative_iff)
have holh: "h holomorphic_on S"
proof (simp add: holomorphic_on_def, clarify)
fix z assume "z ∈ S"
show "h field_differentiable at z within S"
proof (cases "z = ξ")
case True then show ?thesis
using field_differentiable_at_within field_differentiable_def h0 by blast
next
case False
then have "f field_differentiable at z within S"
using holomorphic_onD [OF holf, of z] ‹z ∈ S›
unfolding field_differentiable_def has_field_derivative_iff
by (force intro: exI [where x="dist ξ z"] elim: Lim_transform_within_set [unfolded eventually_at])
then show ?thesis
by (simp add: h_def power2_eq_square derivative_intros)
qed
qed
define g where [abs_def]: "g z = (if z = ξ then deriv h ξ else (h z - h ξ) / (z - ξ))" for z
have §: "∀z∈S - {ξ}. (g z - g ξ) / (z - ξ) = f z"
using h0 by (auto simp: g_def power2_eq_square divide_simps DERIV_imp_deriv h_def)
have "g holomorphic_on S"
unfolding g_def by (rule pole_lemma [OF holh ξ])
then have "(λz. if z = ξ then deriv g ξ else (g z - g ξ) / (z - ξ)) holomorphic_on S"
using ξ pole_lemma by blast
then show ?thesis
using "§" remove_def by fastforce
qed
ultimately show "?P = ?Q" and "?P = ?R"
by meson+
qed
lemma pole_at_infinity:
assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) ⤏ l) at_infinity"
obtains a n where "⋀z. f z = (∑i≤n. a i * z^i)"
proof (cases "l = 0")
case False
show thesis
proof
show "f z = (∑i≤0. inverse l * z ^ i)" for z
using False tendsto_inverse [OF lim] by (simp add: Liouville_weak [OF holf])
qed
next
case True
then have [simp]: "l = 0" .
show ?thesis
proof (cases "∃r. 0 < r ∧ (∀z ∈ ball 0 r - {0}. f(inverse z) ≠ 0)")
case True
then obtain r where "0 < r" and r: "⋀z. z ∈ ball 0 r - {0} ⟹ f(inverse z) ≠ 0"
by auto
have 1: "inverse ∘ f ∘ inverse holomorphic_on ball 0 r - {0}"
by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+
have 2: "0 ∈ interior (ball 0 r)"
using ‹0 < r› by simp
obtain g where holg: "g holomorphic_on ball 0 r"
and geq: "⋀z. z ∈ ball 0 r - {0} ⟹ g z = (inverse ∘ f ∘ inverse) z"
using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] holomorphic_on_extend_bounded [OF 1 2]
by (smt (verit, del_insts) ‹l = 0› eventually_mono norm_conv_dist)
have ifi0: "(inverse ∘ f ∘ inverse) ─0→ 0"
using ‹l = 0› lim lim_at_infinity_0 by blast
have g2g0: "g ─0→ g 0"
using ‹0 < r› centre_in_ball continuous_at continuous_on_eq_continuous_at holg
by (blast intro: holomorphic_on_imp_continuous_on)
have g2g1: "g ─0→ 0"
proof (rule Lim_transform_within_open [OF ifi0 open_ball])
show "⋀x. ⟦x ∈ ball 0 r; x ≠ 0⟧ ⟹ (inverse ∘ f ∘ inverse) x = g x"
by (auto simp: geq)
qed (auto simp: ‹0 < r›)
have [simp]: "g 0 = 0"
by (rule tendsto_unique [OF _ g2g0 g2g1]) simp
have "ball 0 r - {0::complex} ≠ {}"
using ‹0 < r› by (metis "2" Diff_iff insert_Diff interior_ball interior_singleton)
then obtain w::complex where "w ≠ 0" and w: "norm w < r" by force
then have "g w ≠ 0" by (simp add: geq r)
obtain B n e where "0 < B" "0 < e" "e ≤ r"
and leg: "⋀w. norm w < e ⟹ B * cmod w ^ n ≤ cmod (g w)"
proof (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball])
show "g w ≠ g 0"
by (simp add: ‹g w ≠ 0›)
show "w ∈ ball 0 r"
using mem_ball_0 w by blast
qed (use ‹0 < r› in ‹auto simp: ball_subset_ball_iff›)
have §: "cmod (f z) ≤ cmod z ^ n / B" if "2/e ≤ cmod z" for z
proof -
have ize: "inverse z ∈ ball 0 e - {0}" using that ‹0 < e›
by (auto simp: norm_divide field_split_simps algebra_simps)
then have [simp]: "z ≠ 0" and izr: "inverse z ∈ ball 0 r - {0}" using ‹e ≤ r›
by auto
then have [simp]: "f z ≠ 0"
using r [of "inverse z"] by simp
have [simp]: "f z = inverse (g (inverse z))"
using izr geq [of "inverse z"] by simp
show ?thesis using ize leg [of "inverse z"] ‹0 < B› ‹0 < e›
by (simp add: field_split_simps norm_divide algebra_simps)
qed
show thesis
proof
show "f z = (∑i≤n. (deriv ^^ i) f 0 / fact i * z ^ i)" for z
using § by (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp)
qed
next
case False
then have fi0: "⋀r. r > 0 ⟹ ∃z∈ball 0 r - {0}. f (inverse z) = 0"
by simp
have fz0: "f z = 0" if "0 < r" and lt1: "⋀x. x ≠ 0 ⟹ cmod x < r ⟹ inverse (cmod (f (inverse x))) < 1"
for z r
proof -
have f0: "(f ⤏ 0) at_infinity"
proof -
have DIM_complex[intro]: "2 ≤ DIM(complex)"
by simp
have "f (inverse x) ≠ 0 ⟹ x ≠ 0 ⟹ cmod x < r ⟹ 1 < cmod (f (inverse x))" for x
using lt1[of x] by (auto simp: field_simps)
then have **: "cmod (f (inverse x)) ≤ 1 ⟹ x ≠ 0 ⟹ cmod x < r ⟹ f (inverse x) = 0" for x
by force
then have *: "(f ∘ inverse) ` (ball 0 r - {0}) ⊆ {0} ∪ - ball 0 1"
by force
have "continuous_on (inverse ` (ball 0 r - {0})) f"
using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast
then have "connected ((f ∘ inverse) ` (ball 0 r - {0}))"
using connected_punctured_ball
by (intro connected_continuous_image continuous_intros; force)
then have "{0} ∩ (f ∘ inverse) ` (ball 0 r - {0}) = {} ∨ - ball 0 1 ∩ (f ∘ inverse) ` (ball 0 r - {0}) = {}"
by (rule connected_closedD) (use * in auto)
then have "f (inverse w) = 0" if "w ≠ 0" "cmod w < r" for w
using **[of w] fi0 ‹0 < r› that by force
then show ?thesis
unfolding lim_at_infinity_0
using eventually_at ‹r > 0› by (force simp: intro: tendsto_eventually)
qed
obtain w where "w ∈ ball 0 r - {0}" and "f (inverse w) = 0"
using False ‹0 < r› by blast
then show ?thesis
by (auto simp: f0 Liouville_weak [OF holf, of 0])
qed
show thesis
proof
show "⋀z. f z = (∑i≤0. 0 * z ^ i)"
using lim
apply (simp add: lim_at_infinity_0 Lim_at dist_norm norm_inverse)
using fz0 zero_less_one by blast
qed
qed
qed
subsection ‹Entire proper functions are precisely the non-trivial polynomials›
lemma proper_map_polyfun:
fixes c :: "nat ⇒ 'a::{real_normed_div_algebra,heine_borel}"
assumes "closed S" and "compact K" and c: "c i ≠ 0" "1 ≤ i" "i ≤ n"
shows "compact (S ∩ {z. (∑i≤n. c i * z^i) ∈ K})"
proof -
obtain B where "B > 0" and B: "⋀x. x ∈ K ⟹ norm x ≤ B"
by (metis compact_imp_bounded ‹compact K› bounded_pos)
have *: "norm x ≤ b"
if "⋀x. b ≤ norm x ⟹ B + 1 ≤ norm (∑i≤n. c i * x ^ i)"
"(∑i≤n. c i * x ^ i) ∈ K" for b x
proof -
have "norm (∑i≤n. c i * x ^ i) ≤ B"
using B that by blast
moreover have "¬ B + 1 ≤ B"
by simp
ultimately show "norm x ≤ b"
using that by (metis (no_types) less_eq_real_def not_less order_trans)
qed
have "bounded {z. (∑i≤n. c i * z ^ i) ∈ K}"
using Limits.polyfun_extremal [where c=c and B="B+1", OF c]
by (auto simp: bounded_pos eventually_at_infinity_pos *)
moreover have "closed ((λz. (∑i≤n. c i * z ^ i)) -` K)"
using ‹compact K› compact_eq_bounded_closed
by (intro allI continuous_closed_vimage continuous_intros; force)
ultimately show ?thesis
using closed_Int_compact [OF ‹closed S›] compact_eq_bounded_closed
by (auto simp add: vimage_def)
qed
lemma proper_map_polyfun_univ:
fixes c :: "nat ⇒ 'a::{real_normed_div_algebra,heine_borel}"
assumes "compact K" "c i ≠ 0" "1 ≤ i" "i ≤ n"
shows "compact ({z. (∑i≤n. c i * z^i) ∈ K})"
using proper_map_polyfun [of UNIV K c i n] assms by simp
lemma proper_map_polyfun_eq:
assumes "f holomorphic_on UNIV"
shows "(∀k. compact k ⟶ compact {z. f z ∈ k}) ⟷
(∃c n. 0 < n ∧ (c n ≠ 0) ∧ f = (λz. ∑i≤n. c i * z^i))"
(is "?lhs = ?rhs")
proof
assume compf [rule_format]: ?lhs
have 2: "∃k. 0 < k ∧ a k ≠ 0 ∧ f = (λz. ∑i ≤ k. a i * z ^ i)"
if "⋀z. f z = (∑i≤n. a i * z ^ i)" for a n
proof (cases "∀i≤n. 0<i ⟶ a i = 0")
case True
then have [simp]: "⋀z. f z = a 0"
by (simp add: that sum.atMost_shift)
have False using compf [of "{a 0}"] by simp
then show ?thesis ..
next
case False
then obtain k where k: "0 < k" "k≤n" "a k ≠ 0" by force
define m where "m = (GREATEST k. k≤n ∧ a k ≠ 0)"
have m: "m≤n ∧ a m ≠ 0"
unfolding m_def
using GreatestI_nat [where b = n] k by (metis (mono_tags, lifting))
have [simp]: "a i = 0" if "m < i" "i ≤ n" for i
using Greatest_le_nat [where b = "n" and P = "λk. k≤n ∧ a k ≠ 0"]
using m_def not_le that by auto
have "k ≤ m"
unfolding m_def
using Greatest_le_nat [where b = n] k by (metis (mono_tags, lifting))
with k m show ?thesis
by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right)
qed
have *: "((inverse ∘ f) ⤏ 0) at_infinity"
proof (rule Lim_at_infinityI)
fix e::real assume "0 < e"
with compf [of "cball 0 (inverse e)"]
show "∃B. ∀x. B ≤ cmod x ⟶ dist ((inverse ∘ f) x) 0 ≤ e"
apply (clarsimp simp: compact_eq_bounded_closed norm_divide divide_simps mult.commute elim!: bounded_normE_less)
by (meson linorder_not_le nle_le)
qed
then obtain a n where "⋀z. f z = (∑i≤n. a i * z^i)"
using assms pole_at_infinity by blast
with * 2 show ?rhs by blast
next
assume ?rhs
then obtain c n where "0 < n" "c n ≠ 0" "f = (λz. ∑i≤n. c i * z ^ i)" by blast
then have "compact {z. f z ∈ k}" if "compact k" for k
by (auto intro: proper_map_polyfun_univ [OF that])
then show ?lhs by blast
qed
subsection ‹Relating invertibility and nonvanishing of derivative›
lemma has_complex_derivative_locally_injective:
assumes holf: "f holomorphic_on S"
and S: "ξ ∈ S" "open S"
and dnz: "deriv f ξ ≠ 0"
obtains r where "r > 0" "ball ξ r ⊆ S" "inj_on f (ball ξ r)"
proof -
have *: "∃d>0. ∀x. dist ξ x < d ⟶ onorm (λv. deriv f x * v - deriv f ξ * v) < e" if "e > 0" for e
proof -
have contdf: "continuous_on S (deriv f)"
by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on ‹open S›)
obtain δ where "δ>0" and δ: "⋀x. ⟦x ∈ S; dist x ξ ≤ δ⟧ ⟹ cmod (deriv f x - deriv f ξ) ≤ e/2"
using continuous_onE [OF contdf ‹ξ ∈ S›, of "e/2"] ‹0 < e›
by (metis dist_complex_def half_gt_zero less_imp_le)
have §: "⋀ζ z. ⟦ζ ∈ S; dist ξ ζ < δ⟧ ⟹ cmod (deriv f ζ - deriv f ξ) * cmod z ≤ e/2 * cmod z"
by (intro mult_right_mono [OF δ]) (auto simp: dist_commute)
obtain ε where "ε>0" "ball ξ ε ⊆ S"
by (metis openE [OF ‹open S› ‹ξ ∈ S›])
with ‹δ>0› have "∃δ>0. ∀x. dist ξ x < δ ⟶ onorm (λv. deriv f x * v - deriv f ξ * v) ≤ e/2"
using §
apply (rule_tac x="min δ ε" in exI)
apply (intro conjI allI impI Operator_Norm.onorm_le)
apply (force simp: mult_right_mono norm_mult [symmetric] left_diff_distrib δ)+
done
with ‹e>0› show ?thesis by force
qed
have "inj ((*) (deriv f ξ))"
using dnz by simp
then obtain g' where g': "linear g'" "g' ∘ (*) (deriv f ξ) = id"
using linear_injective_left_inverse [of "(*) (deriv f ξ)"] by auto
have fder: "⋀x. x ∈ S ⟹ (f has_derivative (*) (deriv f x)) (at x)"
using ‹open S› has_field_derivative_imp_has_derivative holf holomorphic_derivI by blast
show ?thesis
apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "λz h. deriv f z * h" and g' = g'])
using g' * by (simp_all add: fder linear_conv_bounded_linear that)
qed
lemma has_complex_derivative_locally_invertible:
assumes holf: "f holomorphic_on S"
and S: "ξ ∈ S" "open S"
and dnz: "deriv f ξ ≠ 0"
obtains r where "r > 0" "ball ξ r ⊆ S" "open (f ` (ball ξ r))" "inj_on f (ball ξ r)"
proof -
obtain r where "r > 0" "ball ξ r ⊆ S" "inj_on f (ball ξ r)"
by (blast intro: that has_complex_derivative_locally_injective [OF assms])
then have ξ: "ξ ∈ ball ξ r" by simp
then have nc: "¬ f constant_on ball ξ r"
using ‹inj_on f (ball ξ r)› injective_not_constant by fastforce
have holf': "f holomorphic_on ball ξ r"
using ‹ball ξ r ⊆ S› holf holomorphic_on_subset by blast
have "open (f ` ball ξ r)"
by (simp add: ‹inj_on f (ball ξ r)› holf' open_mapping_thm3)
then show ?thesis
using ‹0 < r› ‹ball ξ r ⊆ S› ‹inj_on f (ball ξ r)› that by blast
qed
lemma holomorphic_injective_imp_regular:
assumes holf: "f holomorphic_on S"
and "open S" and injf: "inj_on f S"
and "ξ ∈ S"
shows "deriv f ξ ≠ 0"
proof -
obtain r where "r>0" and r: "ball ξ r ⊆ S" using assms by (blast elim!: openE)
have holf': "f holomorphic_on ball ξ r"
using ‹ball ξ r ⊆ S› holf holomorphic_on_subset by blast
show ?thesis
proof (cases "∀n>0. (deriv ^^ n) f ξ = 0")
case True
have fcon: "f w = f ξ" if "w ∈ ball ξ r" for w
by (meson open_ball True ‹0 < r› centre_in_ball connected_ball holf'
holomorphic_fun_eq_const_on_connected that)
have False
using fcon [of "ξ + r/2"] ‹0 < r› r injf unfolding inj_on_def
by (metis ‹ξ ∈ S› contra_subsetD dist_commute fcon mem_ball perfect_choose_dist)
then show ?thesis ..
next
case False
then obtain n0 where n0: "n0 > 0 ∧ (deriv ^^ n0) f ξ ≠ 0" by blast
define n where [abs_def]: "n = (LEAST n. n > 0 ∧ (deriv ^^ n) f ξ ≠ 0)"
have n_ne: "n > 0" "(deriv ^^ n) f ξ ≠ 0"
using def_LeastI [OF n_def n0] by auto
have n_min: "⋀k. 0 < k ⟹ k < n ⟹ (deriv ^^ k) f ξ = 0"
using def_Least_le [OF n_def] not_le by auto
obtain g δ where "0 < δ"
and holg: "g holomorphic_on ball ξ δ"
and fd: "⋀w. w ∈ ball ξ δ ⟹ f w - f ξ = ((w - ξ) * g w) ^ n"
and gnz: "⋀w. w ∈ ball ξ δ ⟹ g w ≠ 0"
by (blast intro: n_min holomorphic_factor_order_of_zero_strong [OF holf ‹open S› ‹ξ ∈ S› n_ne])
show ?thesis
proof (cases "n=1")
case True
with n_ne show ?thesis by auto
next
case False
have "g holomorphic_on ball ξ (min r δ)"
using holg by (simp add: holomorphic_on_subset subset_ball)
then have holgw: "(λw. (w - ξ) * g w) holomorphic_on ball ξ (min r δ)"
by (intro holomorphic_intros)
have gd: "⋀w. dist ξ w < δ ⟹ (g has_field_derivative deriv g w) (at w)"
using holg
by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH)
have *: "⋀w. w ∈ ball ξ (min r δ)
⟹ ((λw. (w - ξ) * g w) has_field_derivative ((w - ξ) * deriv g w + g w))
(at w)"
by (rule gd derivative_eq_intros | simp)+
have [simp]: "deriv (λw. (w - ξ) * g w) ξ ≠ 0"
using * [of ξ] ‹0 < δ› ‹0 < r› by (simp add: DERIV_imp_deriv gnz)
obtain T where "ξ ∈ T" "open T" and Tsb: "T ⊆ ball ξ (min r δ)" and oimT: "open ((λw. (w - ξ) * g w) ` T)"
using ‹0 < r› ‹0 < δ› has_complex_derivative_locally_invertible [OF holgw, of ξ]
by force
define U where "U = (λw. (w - ξ) * g w) ` T"
have "open U" by (metis oimT U_def)
moreover have "0 ∈ U"
using ‹ξ ∈ T› by (auto simp: U_def intro: image_eqI [where x = ξ])
ultimately obtain ε where "ε>0" and ε: "cball 0 ε ⊆ U"
using ‹open U› open_contains_cball by blast
then have "ε * exp(2 * of_real pi * 𝗂 * (0/n)) ∈ cball 0 ε"
"ε * exp(2 * of_real pi * 𝗂 * (1/n)) ∈ cball 0 ε"
by (auto simp: norm_mult)
with ε have "ε * exp(2 * of_real pi * 𝗂 * (0/n)) ∈ U"
"ε * exp(2 * of_real pi * 𝗂 * (1/n)) ∈ U" by blast+
then obtain y0 y1 where "y0 ∈ T" and y0: "(y0 - ξ) * g y0 = ε * exp(2 * of_real pi * 𝗂 * (0/n))"
and "y1 ∈ T" and y1: "(y1 - ξ) * g y1 = ε * exp(2 * of_real pi * 𝗂 * (1/n))"
by (auto simp: U_def)
then have "y0 ∈ ball ξ δ" "y1 ∈ ball ξ δ" using Tsb by auto
then have "f y0 - f ξ = ((y0 - ξ) * g y0) ^ n" "f y1 - f ξ = ((y1 - ξ) * g y1) ^ n"
using fd by blast+
moreover have "y0 ≠ y1"
using y0 y1 ‹ε > 0› complex_root_unity_eq_1 [of n 1] ‹n > 0› False by auto
moreover have "T ⊆ S"
by (meson Tsb min.cobounded1 order_trans r subset_ball)
ultimately have False
using inj_onD [OF injf, of y0 y1] ‹y0 ∈ T› ‹y1 ∈ T›
using complex_root_unity [of n 1]
by (auto simp: y0 y1 power_mult_distrib diff_eq_eq n_ne)
then show ?thesis ..
qed
qed
qed
subsubsection ‹Hence a nice clean inverse function theorem›
lemma has_field_derivative_inverse_strong:
fixes f :: "'a::{euclidean_space,real_normed_field} ⇒ 'a"
shows "⟦DERIV f x :> f'; f' ≠ 0; open S; x ∈ S; continuous_on S f;
⋀z. z ∈ S ⟹ g (f z) = z⟧
⟹ DERIV g (f x) :> inverse (f')"
unfolding has_field_derivative_def
by (rule has_derivative_inverse_strong [of S x f g]) auto
lemma has_field_derivative_inverse_strong_x:
fixes f :: "'a::{euclidean_space,real_normed_field} ⇒ 'a"
shows "⟦DERIV f (g y) :> f'; f' ≠ 0; open S; continuous_on S f; g y ∈ S; f(g y) = y;
⋀z. z ∈ S ⟹ g (f z) = z⟧
⟹ DERIV g y :> inverse (f')"
unfolding has_field_derivative_def
by (rule has_derivative_inverse_strong_x [of S g y f]) auto
proposition holomorphic_has_inverse:
assumes holf: "f holomorphic_on S"
and "open S" and injf: "inj_on f S"
obtains g where "g holomorphic_on (f ` S)"
"⋀z. z ∈ S ⟹ deriv f z * deriv g (f z) = 1"
"⋀z. z ∈ S ⟹ g(f z) = z"
proof -
have ofs: "open (f ` S)"
by (rule open_mapping_thm3