diff --git a/src/HOL/Complex_Analysis/Conformal_Mappings.thy b/src/HOL/Complex_Analysis/Conformal_Mappings.thy --- a/src/HOL/Complex_Analysis/Conformal_Mappings.thy +++ b/src/HOL/Complex_Analysis/Conformal_Mappings.thy @@ -1,1961 +1,1894 @@ 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 - apply (rule holomorphic_power_series [OF _ that]) - apply (rule holomorphic_on_subset [OF holf r]) - done + 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" - apply (rule powser_0_nonzero [OF \0 < r\ powf \f \ = 0\ m]) - using \m \ 0\ by (auto simp: dist_commute dist_norm) + 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 + then show thesis apply (rule that) using r s by 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]) - show ?thesis - apply (rule ccontr) - apply (rule isolated_zeros [OF holf \open S\ \connected S\ \\ \ S\ \f \ = 0\ \w \ S\], assumption) + 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 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" by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) 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)" - apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]]) - apply simp - done + 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)" - apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]]) - apply simp - done + 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 \) \ norm (g v)" + then have "cmod (g \) \ cmod (g v)" by simp - with w have wr: "norm (\ - w) = r" and nfw: "norm (f w) \ norm (f \)" - apply (simp_all add: dist_norm) - by (metis \0 < cmod (f \)\ g_def less_imp_inverse_less norm_inverse not_le now order_trans v) + 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 (metis (no_types) \0 < cmod (f \)\ less_imp_inverse_less norm_inverse not_le now order_trans 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" - apply (rule_tac r="s/2" in that) - using \0 < s\ by auto + 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 \))" - apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]]) - apply (simp add: dist_norm) - done + 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 + qed have "continuous_on (cball \ r) (\z. \ - f z)" - apply (rule continuous_intros)+ using \cball \ r \ U\ \f holomorphic_on U\ - apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) - done + by (force intro: continuous_intros continuous_on_subset holomorphic_on_imp_continuous_on) moreover have "(\z. \ - f z) holomorphic_on ball \ r" - apply (rule holomorphic_intros)+ - apply (metis \cball \ r \ U\ \f holomorphic_on U\ holomorphic_on_subset interior_cball interior_subset) - done + 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" - apply (rule holomorphic_contract_to_zero) - apply (blast intro!: \0 < r\ *)+ - done + 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\<^marker>\tag unimportant\ 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" - apply (rule fnc) - using C \C \ S\ \C \ components S\ in_components_nonempty by auto + 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))" - apply (rule open_mapping_thm [OF _ C _ _ nf]) - apply (simp add: C \open U\ open_Int, blast) - done + by (meson C \open U\ inf_le1 nf open_Int open_mapping_thm) } ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ open_mapping_thm3: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" shows "open (f ` S)" -apply (rule open_mapping_thm2 [OF holf]) -using assms -apply (simp_all add:) -using injective_not_constant subset_inj_on by blast +proof (rule open_mapping_thm2 [OF holf]) + show "\X. \open X; X \ S; X \ {}\ \ \ f constant_on X" + using inj_on_subset injective_not_constant injf by blast +qed (use assms in auto) 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) - using that 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 zin: "z \ closure S" and z: "\y. y \ closure S \ (cmod \ f) y \ (cmod \ f) z" + 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 zin: "z \ connected_component_set (interior S) z" - by (simp add: 2) have "f constant_on (connected_component_set (interior S) z)" - apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin]) - apply (metis connected_component_subset holf holomorphic_on_subset) - apply (simp_all add: open_connected_component) - by (metis closure_subset comp_eq_dest_lhs interior_subset subsetCE z connected_component_in) + 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}" - apply (rule image_closure_subset) - apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset) - using c - apply auto - done + 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 "frontier(connected_component_set (interior S) z) \ {}" - apply (simp add: frontier_eq_empty) - by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV) + 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" - apply (rule leB) - using w -using frontier_interior_subset frontier_of_connected_component_subset by blast + 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\<^marker>\tag unimportant\ 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\<^marker>\tag unimportant\ \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: "{.. = 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" - apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric]) - apply simp - apply (simp only: dfz sing) - apply (simp add: powf_def) - done + 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" - apply (rule sums_unique2) - apply (rule fsums [unfolded powf_def]) using sums_mult [OF sumsg, of "(w - \) ^ n"] - by (auto simp: power_add mult_ac) + 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" - apply (rule exE [OF continuous_on_avoid [OF contg _ \g \ \ 0\]]) - using \0 < r\ - apply force - by (metis \0 < r\ less_trans mem_ball not_less_iff_gr_or_eq) + using \0 < r\ continuous_on_avoid [OF contg _ \g \ \ 0\] + by (metis centre_in_ball le_cases mem_ball mem_ball_leI) show ?thesis - apply (rule that [where g=g and r ="min r d"]) - using \0 < r\ \0 < d\ holg - apply (auto simp: feq holomorphic_on_subset subset_ball d) - done + 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: "\x. dist \ x < r \ (\z. deriv g z / g z) field_differentiable at x" - apply (rule derivative_intros)+ - using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) - apply (metis open_ball at_within_open holg holomorphic_on_def mem_ball) - using gne mem_ball by blast - obtain h where h: "\x. x \ ball \ r \ (h has_field_derivative deriv g x / g x) (at x)" - apply (rule exE [OF holomorphic_convex_primitive [of "ball \ r" "{}" "\z. deriv g z / g z"]]) - apply (auto simp: con cd) - apply (metis open_ball at_within_open mem_ball) - done + 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 0: "dist \ x < r \ ((\x. exp (h x) / g x) has_field_derivative 0) (at x)" for x - apply (rule h derivative_eq_intros | simp)+ - apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2]) - using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h) - done + apply (rule h derivative_eq_intros DERIV_deriv_iff_field_differentiable [THEN iffD2] | simp)+ + using holg by (auto simp: holomorphic_on_imp_differentiable_at gne h) 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" - apply (rule holomorphic_on_compose [unfolded o_def, where g = exp]) - apply (rule holomorphic_intros)+ - using h holomorphic_on_open apply blast - apply (rule holomorphic_intros)+ - using \0 < n\ apply simp - apply (rule holomorphic_intros)+ - done + 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 - apply (rule that [where g="\z. exp((Ln(inverse c) + h z)/n)" and r =r]) - using \0 < r\ \0 < n\ - apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric]) - apply (rule hol) - apply (simp add: Transcendental.exp_add gne) - done + 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 "0 < r1" "g holomorphic_on ball \ r1" - "\w. w \ ball \ r1 \ f w = (w - \) ^ n * g w" - "\w. w \ ball \ r1 \ g w \ 0" + 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\) - then show ?thesis - apply (rule_tac g=g and r="min r0 r1" and n=n in that) - using \0 < n\ \0 < r0\ \0 < r1\ \ball \ r0 \ S\ - apply (auto simp: subset_ball intro: holomorphic_on_subset) - done + 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" . - moreover have "x \ 0" using gnz x \d < r\ by auto - ultimately show ?thesis - apply (rule_tac k="norm x" and n=n and r=d in that) - using \d < r\ leg - apply (auto simp: \0 < d\ fne norm_mult norm_power algebra_simps mult_right_mono) - done + 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 *: "\\<^sub>F z in at \. dist (g z) (g \) < 1 \ cmod (f z) \ cmod (g \) + 1" - apply (simp add: eventually_at) - apply (rule_tac x="\" in exI) - using \ \0 < \\ - apply (clarsimp simp:) - apply (drule_tac c=x in subsetD) - apply (simp add: dist_commute) - by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD) + 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 *: "\\<^sub>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 show ?thesis apply (rule_tac x="norm(g \) + 1" in exI) apply (rule eventually_mp [OF * tendstoD [where e=1]], auto) done qed moreover have "?Q" if "\\<^sub>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 h0: "(h has_field_derivative 0) (at \)" apply (simp add: h_def has_field_derivative_iff) - apply (rule Lim_transform_within [OF that, of 1]) - apply (auto simp: field_split_simps power2_eq_square) + apply (auto simp: field_split_simps power2_eq_square Lim_transform_within [OF that, of 1]) done 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 holg: "g holomorphic_on S" unfolding g_def by (rule pole_lemma [OF holh \]) + 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) show ?thesis - apply (rule_tac x="\z. if z = \ then deriv g \ else (g z - g \)/(z - \)" in exI) - apply (rule conjI) - apply (rule pole_lemma [OF holg \]) - apply (auto simp: g_def power2_eq_square divide_simps) - using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square) + apply (intro exI conjI) + apply (rule pole_lemma [OF holg \]) + apply (simp add: \
) done 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 - with tendsto_inverse [OF lim] show ?thesis - apply (rule_tac a="(\n. inverse l)" and n=0 in that) - apply (simp add: Liouville_weak [OF holf, of "inverse l"]) - done + 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 have "\B. 0 eventually (\z. cmod ((inverse \ f \ inverse) z) \ B) (at 0)" apply (rule exI [where x=1]) - apply simp using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] - apply (rule eventually_mono) - apply (simp add: dist_norm) - done + by (simp add: eventually_mono) with holomorphic_on_extend_bounded [OF 1 2] obtain g where holg: "g holomorphic_on ball 0 r" and geq: "\z. z \ ball 0 r - {0} \ g z = (inverse \ f \ inverse) z" by meson 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" apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]]) using \0 < r\ by (auto simp: geq) have [simp]: "g 0 = 0" by (rule tendsto_unique [OF _ g2g0 g2g1]) simp have "ball 0 r - {0::complex} \ {}" - using \0 < r\ - apply (clarsimp simp: ball_def dist_norm) - apply (drule_tac c="of_real r/2" in subsetD, auto) - done + 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)" - apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w]) - using \0 < r\ w \g w \ 0\ by (auto simp: ball_subset_ball_iff) - have "cmod (f z) \ cmod z ^ n / B" if "2/e \ cmod z" for z + 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 - then show ?thesis - apply (rule_tac a = "\k. (deriv ^^ k) f 0 / (fact k)" and n=n in that) - apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp) - done + 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)" \ \should not be necessary!\ 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}))" - apply (intro connected_continuous_image continuous_intros) - apply (force intro: connected_punctured_ball)+ - done + 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 "w \ 0 \ cmod w < r \ f (inverse w) = 0" for w - using fi0 **[of w] \0 < r\ - apply (auto simp add: inf.commute [of "- ball 0 1"] Diff_eq [symmetric] image_subset_iff dest: less_imp_le) - apply fastforce - apply (drule bspec [of _ _ w]) - apply (auto dest: less_imp_le) - done + 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 - apply (simp add: lim_at_infinity_0) - apply (rule tendsto_eventually) - apply (simp add: eventually_at) - apply (rule_tac x=r in exI) - apply (simp add: \0 < r\ dist_norm) - done + unfolding lim_at_infinity_0 + using eventually_at \r > 0\ by (force simp add: 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 - apply (rule that [of "\n. 0" 0]) - using lim [unfolded lim_at_infinity_0] - apply (simp add: Lim_at dist_norm norm_inverse) - apply (drule_tac x=1 in spec) - using fz0 apply auto - done + 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\<^marker>\tag unimportant\ \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)" - apply (intro allI continuous_closed_vimage continuous_intros) - using \compact K\ compact_eq_bounded_closed by blast + 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 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 - apply (rule GreatestI_nat [where b = n]) - using k apply auto - done + 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 - apply (rule Greatest_le_nat [where b = "n"]) - using k apply auto - done + 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" + 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 simp apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse) - apply (rule_tac x="b+1" in exI) - 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) - done + by (metis (no_types, hide_lams) inverse_inverse_eq le_less_trans less_eq_real_def less_imp_inverse_less linordered_field_no_ub not_less) qed - then show ?rhs - apply (rule pole_at_infinity [OF assms]) - using 2 apply blast - done + 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 simp - apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult) - apply (rule mult_right_mono [OF \]) - apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \) + 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 simp: linear_times) + using linear_injective_left_inverse [of "(*) (deriv f \)"] by auto 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' * + using g' * apply (simp_all add: linear_conv_bounded_linear that) - using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf - holomorphic_on_imp_differentiable_at \open S\ apply blast - done + using \open S\ has_field_derivative_imp_has_derivative holf holomorphic_derivI by blast 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)" - apply (rule open_mapping_thm [OF holf']) - using nc apply auto - done + 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 - apply (rule holomorphic_fun_eq_const_on_connected [OF holf']) - using True \0 < r\ that by auto + 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" - apply (rule holomorphic_factor_order_of_zero_strong [OF holf \open S\ \\ \ S\ n_ne]) - apply (blast intro: n_min)+ - done + 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 holgw: "(\w. (w - \) * g w) holomorphic_on ball \ (min r \)" - apply (rule holomorphic_intros)+ + 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)" - apply (rule has_complex_derivative_locally_invertible [OF holgw, of \]) - using \0 < r\ \0 < \\ - apply (simp_all add:) - by (meson open_ball centre_in_ball) + 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 fd [of y0] fd [of y1] complex_root_unity [of n 1] n_ne + using complex_root_unity [of n 1] apply (simp add: y0 y1 power_mult_distrib) apply (force simp: algebra_simps) done then show ?thesis .. qed qed qed text\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 [OF assms]) have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \ S" for z proof - have 1: "(f has_field_derivative deriv f z) (at z)" using DERIV_deriv_iff_field_differentiable \z \ S\ \open S\ holf holomorphic_on_imp_differentiable_at by blast have 2: "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast show ?thesis - apply (rule has_field_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) - apply (simp add: holf holomorphic_on_imp_continuous_on) - by (simp add: injf the_inv_into_f_f) + proof (rule has_field_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) + show "continuous_on S f" + by (simp add: holf holomorphic_on_imp_continuous_on) + show "\z. z \ S \ the_inv_into S f (f z) = z" + by (simp add: injf the_inv_into_f_f) + qed qed show ?thesis proof show "the_inv_into S f holomorphic_on f ` S" by (simp add: holomorphic_on_open ofs) (blast intro: *) next fix z assume "z \ S" have "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast then show "deriv f z * deriv (the_inv_into S f) (f z) = 1" using * [OF \z \ S\] by (simp add: DERIV_imp_deriv) next fix z assume "z \ S" show "the_inv_into S f (f z) = z" by (simp add: \z \ S\ injf the_inv_into_f_f) qed qed subsection\The Schwarz Lemma\ lemma Schwarz1: assumes holf: "f holomorphic_on S" and contf: "continuous_on (closure S) f" and S: "open S" "connected S" and boS: "bounded S" and "S \ {}" obtains w where "w \ frontier S" "\z. z \ closure S \ norm (f z) \ norm (f w)" proof - have connf: "continuous_on (closure S) (norm o f)" using contf continuous_on_compose continuous_on_norm_id by blast have coc: "compact (closure S)" by (simp add: \bounded S\ bounded_closure compact_eq_bounded_closed) then obtain x where x: "x \ closure S" and xmax: "\z. z \ closure S \ norm(f z) \ norm(f x)" - apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]]) - using \S \ {}\ apply auto - done + using continuous_attains_sup [OF _ _ connf] \S \ {}\ by auto then show ?thesis proof (cases "x \ frontier S") case True then show ?thesis using that xmax by blast next case False then have "x \ S" using \open S\ frontier_def interior_eq x by auto then have "f constant_on S" - apply (rule maximum_modulus_principle [OF holf S \open S\ order_refl]) - using closure_subset apply (blast intro: xmax) - done + proof (rule maximum_modulus_principle [OF holf S \open S\ order_refl]) + show "\z. z \ S \ cmod (f z) \ cmod (f x)" + using closure_subset by (blast intro: xmax) + qed then have "f constant_on (closure S)" by (rule constant_on_closureI [OF _ contf]) then obtain c where c: "\x. x \ closure S \ f x = c" by (meson constant_on_def) obtain w where "w \ frontier S" by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV) then show ?thesis by (simp add: c frontier_def that) qed qed lemma Schwarz2: "\f holomorphic_on ball 0 r; 0 < s; ball w s \ ball 0 r; \z. norm (w-z) < s \ norm(f z) \ norm(f w)\ \ f constant_on ball 0 r" by (rule maximum_modulus_principle [where U = "ball w s" and \ = w]) (simp_all add: dist_norm) lemma Schwarz3: assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" obtains h where "h holomorphic_on (ball 0 r)" and "\z. norm z < r \ f z = z * (h z)" and "deriv f 0 = h 0" proof - define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z have d0: "deriv f 0 = h 0" by (simp add: h_def) moreover have "h holomorphic_on (ball 0 r)" by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) moreover have "norm z < r \ f z = z * h z" for z by (simp add: h_def) ultimately show ?thesis using that by blast qed proposition Schwarz_Lemma: assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" and no: "\z. norm z < 1 \ norm (f z) < 1" and \: "norm \ < 1" shows "norm (f \) \ norm \" and "norm(deriv f 0) \ 1" and "((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) \ \\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1" (is "?P \ ?Q") proof - obtain h where holh: "h holomorphic_on (ball 0 1)" and fz_eq: "\z. norm z < 1 \ f z = z * (h z)" and df0: "deriv f 0 = h 0" by (rule Schwarz3 [OF holf]) auto have noh_le: "norm (h z) \ 1" if z: "norm z < 1" for z proof - have "norm (h z) < a" if a: "1 < a" for a proof - have "max (inverse a) (norm z) < 1" using z a by (simp_all add: inverse_less_1_iff) then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" using Rats_dense_in_real by blast then have nzr: "norm z < r" and ira: "inverse r < a" using z a less_imp_inverse_less by force+ then have "0 < r" by (meson norm_not_less_zero not_le order.strict_trans2) have holh': "h holomorphic_on ball 0 r" by (meson holh \r < 1\ holomorphic_on_subset less_eq_real_def subset_ball) have conth': "continuous_on (cball 0 r) h" by (meson \r < 1\ dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) obtain w where w: "norm w = r" and lenw: "\z. norm z < r \ norm(h z) \ norm(h w)" apply (rule Schwarz1 [OF holh']) using conth' \0 < r\ by auto have "h w = f w / w" using fz_eq \r < 1\ nzr w by auto then have "cmod (h z) < inverse r" by (metis \0 < r\ \r < 1\ divide_strict_right_mono inverse_eq_divide le_less_trans lenw no norm_divide nzr w) then show ?thesis using ira by linarith qed then show "norm (h z) \ 1" using not_le by blast qed show "cmod (f \) \ cmod \" proof (cases "\ = 0") case True then show ?thesis by auto next case False then show ?thesis by (simp add: noh_le fz_eq \ mult_left_le norm_mult) qed show no_df0: "norm(deriv f 0) \ 1" by (simp add: \\z. cmod z < 1 \ cmod (h z) \ 1\ df0) show "?Q" if "?P" using that proof assume "\z. cmod z < 1 \ z \ 0 \ cmod (f z) = cmod z" then obtain \ where \: "cmod \ < 1" "\ \ 0" "cmod (f \) = cmod \" by blast then have [simp]: "norm (h \) = 1" by (simp add: fz_eq norm_mult) - have "ball \ (1 - cmod \) \ ball 0 1" + have \
: "ball \ (1 - cmod \) \ ball 0 1" by (simp add: ball_subset_ball_iff) moreover have "\z. cmod (\ - z) < 1 - cmod \ \ cmod (h z) \ cmod (h \)" - apply (simp add: algebra_simps) - by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4) + by (metis \cmod (h \) = 1\ \
dist_0_norm dist_complex_def in_mono mem_ball noh_le) ultimately obtain c where c: "\z. norm z < 1 \ h z = c" using Schwarz2 [OF holh, of "1 - norm \" \, unfolded constant_on_def] \ by auto then have "norm c = 1" using \ by force with c show ?thesis using fz_eq by auto next assume [simp]: "cmod (deriv f 0) = 1" then obtain c where c: "\z. norm z < 1 \ h z = c" using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le by auto moreover have "norm c = 1" using df0 c by auto ultimately show ?thesis using fz_eq by auto qed qed corollary Schwarz_Lemma': assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" and no: "\z. norm z < 1 \ norm (f z) < 1" shows "((\\. norm \ < 1 \ norm (f \) \ norm \) \ norm(deriv f 0) \ 1) \ (((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) \ (\\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1))" using Schwarz_Lemma [OF assms] by (metis (no_types) norm_eq_zero zero_less_one) subsection\The Schwarz reflection principle\ lemma hol_pal_lem0: assumes "d \ a \ k" "k \ d \ b" obtains c where "c \ closed_segment a b" "d \ c = k" "\z. z \ closed_segment a c \ d \ z \ k" "\z. z \ closed_segment c b \ k \ d \ z" proof - obtain c where cin: "c \ closed_segment a b" and keq: "k = d \ c" using connected_ivt_hyperplane [of "closed_segment a b" a b d k] by (auto simp: assms) have "closed_segment a c \ {z. d \ z \ k}" "closed_segment c b \ {z. k \ d \ z}" unfolding segment_convex_hull using assms keq by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal) then show ?thesis using cin that by fastforce qed lemma hol_pal_lem1: assumes "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" "d \ 0" and lek: "d \ a \ k" "d \ b \ k" "d \ c \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof - have "interior (convex hull {a, b, c}) \ interior(S \ {x. d \ x \ k})" - apply (rule interior_mono) - apply (rule hull_minimal) - apply (simp add: abc lek) - apply (rule convex_Int [OF \convex S\ convex_halfspace_le]) - done + proof (intro interior_mono hull_minimal) + show "{a, b, c} \ S \ {x. d \ x \ k}" + by (simp add: abc lek) + show "convex (S \ {x. d \ x \ k})" + by (rule convex_Int [OF \convex S\ convex_halfspace_le]) + qed also have "... \ {z \ S. d \ z < k}" by (force simp: interior_open [OF \open S\] \d \ 0\) finally have *: "interior (convex hull {a, b, c}) \ {z \ S. d \ z < k}" . have "continuous_on (convex hull {a,b,c}) f" using \convex S\ contf abc continuous_on_subset subset_hull by fastforce moreover have "f holomorphic_on interior (convex hull {a,b,c})" by (rule holomorphic_on_subset [OF holf1 *]) ultimately show ?thesis using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3 by blast qed lemma hol_pal_lem2: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and lek: "d \ a \ k" "d \ b \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ c \ k") case True show ?thesis by (rule hol_pal_lem1 [OF S abc \d \ 0\ lek True holf1 contf]) next case False then have "d \ c > k" by force obtain a' where a': "a' \ closed_segment b c" and "d \ a' = k" and ba': "\z. z \ closed_segment b a' \ d \ z \ k" and a'c: "\z. z \ closed_segment a' c \ k \ d \ z" - apply (rule hol_pal_lem0 [of d b k c, OF \d \ b \ k\]) - using False by auto + using False hol_pal_lem0 [of d b k c, OF \d \ b \ k\] by auto obtain b' where b': "b' \ closed_segment a c" and "d \ b' = k" and ab': "\z. z \ closed_segment a b' \ d \ z \ k" and b'c: "\z. z \ closed_segment b' c \ k \ d \ z" - apply (rule hol_pal_lem0 [of d a k c, OF \d \ a \ k\]) - using False by auto + using False hol_pal_lem0 [of d a k c, OF \d \ a \ k\] by auto have a'b': "a' \ S \ b' \ S" using a' abc b' convex_contains_segment \convex S\ by auto have "continuous_on (closed_segment c a) f" by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) then have 1: "contour_integral (linepath c a) f = contour_integral (linepath c b') f + contour_integral (linepath b' a) f" - apply (rule contour_integral_split_linepath) - using b' by (simp add: closed_segment_commute) + using b' closed_segment_commute contour_integral_split_linepath by blast have "continuous_on (closed_segment b c) f" by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) then have 2: "contour_integral (linepath b c) f = contour_integral (linepath b a') f + contour_integral (linepath a' c) f" by (rule contour_integral_split_linepath [OF _ a']) have 3: "contour_integral (reversepath (linepath b' a')) f = - contour_integral (linepath b' a') f" by (rule contour_integral_reversepath [OF valid_path_linepath]) have fcd_le: "f field_differentiable at x" if "x \ interior S \ x \ interior {x. d \ x \ k}" for x proof - have "f holomorphic_on S \ {c. d \ c < k}" by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1) then have "\C D. x \ interior C \ interior D \ f holomorphic_on interior C \ interior D" using that by (metis Collect_mem_eq Int_Collect \d \ 0\ interior_halfspace_le interior_open \open S\) then show "f field_differentiable at x" by (metis at_within_interior holomorphic_on_def interior_Int interior_interior) qed have ab_le: "\x. x \ closed_segment a b \ d \ x \ k" proof - fix x :: complex assume "x \ closed_segment a b" then have "\C. x \ C \ b \ C \ a \ C \ \ convex C" by (meson contra_subsetD convex_contains_segment) then show "d \ x \ k" by (metis lek convex_halfspace_le mem_Collect_eq) qed + have cs: "closed_segment a' b' \ {x. d \ x \ k} \ closed_segment b' a \ {x. d \ x \ k}" + by (simp add: \d \ a' = k\ \d \ b' = k\ closed_segment_subset convex_halfspace_le lek(1)) have "continuous_on (S \ {x. d \ x \ k}) f" using contf by (simp add: continuous_on_subset) then have "(f has_contour_integral 0) (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)" apply (rule Cauchy_theorem_convex [where K = "{}"]) - apply (simp_all add: path_image_join convex_Int convex_halfspace_le \convex S\ fcd_le ab_le - closed_segment_subset abc a'b' ba') - by (metis \d \ a' = k\ \d \ b' = k\ convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl) + by (simp_all add: path_image_join convex_Int convex_halfspace_le \convex S\ fcd_le ab_le + closed_segment_subset abc a'b' ba' cs) then have 4: "contour_integral (linepath a b) f + contour_integral (linepath b a') f + contour_integral (linepath a' b') f + contour_integral (linepath b' a) f = 0" by (rule has_chain_integral_chain_integral4) have fcd_ge: "f field_differentiable at x" if "x \ interior S \ x \ interior {x. k \ d \ x}" for x proof - have f2: "f holomorphic_on S \ {c. k < d \ c}" by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2) have f3: "interior S = S" by (simp add: interior_open \open S\) then have "x \ S \ interior {c. k \ d \ c}" using that by simp then show "f field_differentiable at x" using f3 f2 unfolding holomorphic_on_def by (metis (no_types) \d \ 0\ at_within_interior interior_Int interior_halfspace_ge interior_interior) qed + have cs: "closed_segment c b' \ {x. k \ d \ x} \ closed_segment b' a' \ {x. k \ d \ x}" + by (simp add: \d \ a' = k\ b'c closed_segment_subset convex_halfspace_ge) have "continuous_on (S \ {x. k \ d \ x}) f" using contf by (simp add: continuous_on_subset) then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')" apply (rule Cauchy_theorem_convex [where K = "{}"]) - apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \convex S\ - fcd_ge closed_segment_subset abc a'b' a'c) - by (metis \d \ a' = k\ b'c closed_segment_commute convex_contains_segment - convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl) + by (simp_all add: path_image_join convex_Int convex_halfspace_ge \convex S\ + fcd_ge closed_segment_subset abc a'b' a'c cs) then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0" by (rule has_chain_integral_chain_integral3) show ?thesis using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath) qed lemma hol_pal_lem3: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and lek: "d \ a \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ b \ k") case True show ?thesis by (rule hol_pal_lem2 [OF S abc \d \ 0\ lek True holf1 holf2 contf]) next case False show ?thesis proof (cases "d \ c \ k") case True have "contour_integral (linepath c a) f + contour_integral (linepath a b) f + contour_integral (linepath b c) f = 0" by (rule hol_pal_lem2 [OF S \c \ S\ \a \ S\ \b \ S\ \d \ 0\ \d \ c \ k\ lek holf1 holf2 contf]) then show ?thesis by (simp add: algebra_simps) next case False have "contour_integral (linepath b c) f + contour_integral (linepath c a) f + contour_integral (linepath a b) f = 0" - apply (rule hol_pal_lem2 [OF S \b \ S\ \c \ S\ \a \ S\, of "-d" "-k"]) + using hol_pal_lem2 [OF S \b \ S\ \c \ S\ \a \ S\, of "-d" "-k"] using \d \ 0\ \\ d \ b \ k\ False by (simp_all add: holf1 holf2 contf) then show ?thesis by (simp add: algebra_simps) qed qed lemma hol_pal_lem4: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ a \ k") case True show ?thesis by (rule hol_pal_lem3 [OF S abc \d \ 0\ True holf1 holf2 contf]) next case False show ?thesis - apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"]) - using \d \ 0\ False by (simp_all add: holf1 holf2 contf) + using \d \ 0\ hol_pal_lem3 [OF S abc, of "-d" "-k"] False + by (simp_all add: holf1 holf2 contf) qed lemma holomorphic_on_paste_across_line: assumes S: "open S" and "d \ 0" and holf1: "f holomorphic_on (S \ {z. d \ z < k})" and holf2: "f holomorphic_on (S \ {z. k < d \ z})" and contf: "continuous_on S f" shows "f holomorphic_on S" proof - have *: "\t. open t \ p \ t \ continuous_on t f \ (\a b c. convex hull {a, b, c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" if "p \ S" for p proof - obtain e where "e>0" and e: "ball p e \ S" using \p \ S\ openE S by blast then have "continuous_on (ball p e) f" using contf continuous_on_subset by blast moreover have "f holomorphic_on {z. dist p z < e \ d \ z < k}" apply (rule holomorphic_on_subset [OF holf1]) using e by auto moreover have "f holomorphic_on {z. dist p z < e \ k < d \ z}" apply (rule holomorphic_on_subset [OF holf2]) using e by auto ultimately show ?thesis apply (rule_tac x="ball p e" in exI) - using \e > 0\ e \d \ 0\ - apply (simp add:, clarify) - apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k]) - apply (auto simp: subset_hull) - done + using \e > 0\ e \d \ 0\ hol_pal_lem4 [of "ball p e" _ _ _ d _ k] + by (force simp add: subset_hull) qed show ?thesis by (blast intro: * Morera_local_triangle analytic_imp_holomorphic) qed proposition Schwarz_reflection: assumes "open S" and cnjs: "cnj ` S \ S" and holf: "f holomorphic_on (S \ {z. 0 < Im z})" and contf: "continuous_on (S \ {z. 0 \ Im z}) f" and f: "\z. \z \ S; z \ \\ \ (f z) \ \" shows "(\z. if 0 \ Im z then f z else cnj(f(cnj z))) holomorphic_on S" proof - have 1: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. 0 < Im z})" by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf]) have cont_cfc: "continuous_on (S \ {z. Im z \ 0}) (cnj o f o cnj)" - apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) - using cnjs apply auto - done + using cnjs + by (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) auto have "cnj \ f \ cnj field_differentiable at x within S \ {z. Im z < 0}" if "x \ S" "Im x < 0" "f field_differentiable at (cnj x) within S \ {z. 0 < Im z}" for x using that - apply (simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm, clarify) + apply (clarsimp simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm) apply (rule_tac x="cnj f'" in exI) apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify) apply (drule_tac x="cnj xa" in bspec) using cnjs apply force apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj) done then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \ {z. Im z < 0})" using holf cnjs by (force simp: holomorphic_on_def) have 2: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. Im z < 0})" apply (rule iffD1 [OF holomorphic_cong [OF refl]]) using hol_cfc by auto have [simp]: "(S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0}) = S" by force + have eq: "\z. \z \ S; Im z \ 0; 0 \ Im z\ \ f z = cnj (f (cnj z))" + using f Reals_cnj_iff complex_is_Real_iff by auto have "continuous_on ((S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0})) (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" apply (rule continuous_on_cases_local) using cont_cfc contf - apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge) - using f Reals_cnj_iff complex_is_Real_iff apply auto - done + by (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge eq) then have 3: "continuous_on S (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" by force show ?thesis apply (rule holomorphic_on_paste_across_line [OF \open S\, of "- \" _ 0]) - using 1 2 3 - apply auto - done + using 1 2 3 by auto qed subsection\Bloch's theorem\ lemma Bloch_lemma_0: assumes holf: "f holomorphic_on cball 0 r" and "0 < r" and [simp]: "f 0 = 0" and le: "\z. norm z < r \ norm(deriv f z) \ 2 * norm(deriv f 0)" shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \ f ` ball 0 r" proof - have "sqrt 2 < 3/2" by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp show ?thesis proof (cases "deriv f 0 = 0") case True then show ?thesis by simp next case False define C where "C = 2 * norm(deriv f 0)" have "0 < C" using False by (simp add: C_def) have holf': "f holomorphic_on ball 0 r" using holf using ball_subset_cball holomorphic_on_subset by blast then have holdf': "deriv f holomorphic_on ball 0 r" by (rule holomorphic_deriv [OF _ open_ball]) have "Le1": "norm(deriv f z - deriv f 0) \ norm z / (r - norm z) * C" if "norm z < r" for z proof - have T1: "norm(deriv f z - deriv f 0) \ norm z / (R - norm z) * C" if R: "norm z < R" "R < r" for R proof - have "0 < R" using R by (metis less_trans norm_zero zero_less_norm_iff) have df_le: "\x. norm x < r \ norm (deriv f x) \ C" using le by (simp add: C_def) have hol_df: "deriv f holomorphic_on cball 0 R" - apply (rule holomorphic_on_subset) using R holdf' by auto + using R holdf' holomorphic_on_subset by auto have *: "((\w. deriv f w / (w - z)) has_contour_integral 2 * pi * \ * deriv f z) (circlepath 0 R)" if "norm z < R" for z using \0 < R\ that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"] by (force simp: winding_number_circlepath) have **: "((\x. deriv f x / (x - z) - deriv f x / x) has_contour_integral of_real (2 * pi) * \ * (deriv f z - deriv f 0)) (circlepath 0 R)" using has_contour_integral_diff [OF * [of z] * [of 0]] \0 < R\ that by (simp add: algebra_simps) have [simp]: "\x. norm x = R \ x \ z" using that(1) by blast have "norm (deriv f x / (x - z) - deriv f x / x) \ C * norm z / (R * (R - norm z))" if "norm x = R" for x proof - have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) = norm (deriv f x) * norm z" by (simp add: norm_mult right_diff_distrib') show ?thesis - using \0 < R\ \0 < C\ R that - apply (simp add: norm_mult norm_divide divide_simps) - using df_le norm_triangle_ineq2 \0 < C\ apply (auto intro!: mult_mono) - done + using \0 < R\ \0 < C\ R that + by (auto simp add: norm_mult norm_divide divide_simps df_le mult_mono norm_triangle_ineq2) qed then show ?thesis using has_contour_integral_bound_circlepath [OF **, of "C * norm z/(R*(R - norm z))"] \0 < R\ \0 < C\ R apply (simp add: norm_mult norm_divide) apply (simp add: divide_simps mult.commute) done qed obtain r' where r': "norm z < r'" "r' < r" using Rats_dense_in_real [of "norm z" r] \norm z < r\ by blast then have [simp]: "closure {r'<.. norm(f z)" if r: "norm z < r" for z proof - have 1: "\x. x \ ball 0 r \ ((\z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0) (at x within ball 0 r)" by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+ have 2: "closed_segment 0 z \ ball 0 r" by (metis \0 < r\ convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that) - have 3: "(\t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" - apply (rule integrable_on_cmult_right [where 'b=real, simplified]) - apply (rule integrable_on_cdivide [where 'b=real, simplified]) - apply (rule integrable_on_cmult_left [where 'b=real, simplified]) - apply (rule ident_integrable_on) - done have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm z * norm z * x * C / (r - norm z)" if x: "0 \ x" "x \ 1" for x proof - have [simp]: "x * norm z < r" using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero) have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \ norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C" apply (rule Le1) using r x \0 < r\ by simp also have "... \ norm (x *\<^sub>R z) / (r - norm z) * C" using r x \0 < r\ apply (simp add: field_split_simps) by (simp add: \0 < C\ mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono) finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm (x *\<^sub>R z) / (r - norm z) * C * norm z" by (rule mult_right_mono) simp with x show ?thesis by (simp add: algebra_simps) qed have le_norm: "abc \ norm d - e \ norm(f - d) \ e \ abc \ norm f" for abc d e and f::complex by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans) have "norm (integral {0..1} (\x. (deriv f (x *\<^sub>R z) - deriv f 0) * z)) \ integral {0..1} (\t. (norm z)\<^sup>2 * t / (r - norm z) * C)" - apply (rule integral_norm_bound_integral) - using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 - apply (simp add: has_contour_integral_linepath has_integral_integrable_integral) - apply (rule 3) - apply (simp add: norm_mult power2_eq_square 4) - done + proof (rule integral_norm_bound_integral) + show "(\x. (deriv f (x *\<^sub>R z) - deriv f 0) * z) integrable_on {0..1}" + using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 + by (simp add: has_contour_integral_linepath has_integral_integrable_integral) + have "(*) ((cmod z)\<^sup>2) integrable_on {0..1}" + by (metis ident_integrable_on integrable_0 integrable_eq integrable_on_cmult_iff lambda_zero) + then show "(\t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" + using integrable_on_cmult_right[where 'b=real, simplified] integrable_on_cdivide [where 'b=real, simplified] + by blast + qed (simp add: norm_mult power2_eq_square 4) then have int_le: "norm (f z - deriv f 0 * z) \ (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))" using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 - apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) - done + by (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) + have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \ norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" + by (simp add: algebra_simps) + then have \
: "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \ norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" + by (simp add: algebra_simps) show ?thesis apply (rule le_norm [OF _ int_le]) using \norm z < r\ - apply (simp add: power2_eq_square divide_simps C_def norm_mult) - proof - - have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \ norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" - by (simp add: algebra_simps) - then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \ norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" - by (simp add: algebra_simps) - qed + by (simp add: power2_eq_square divide_simps C_def norm_mult \
) qed have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2) < 1" by (auto simp: sqrt2_less_2) have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f" - apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) - apply (subst closure_ball) - using \0 < r\ mult_pos_pos sq201 - apply (auto simp: cball_subset_cball_iff) - done + proof (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) + show "closure (ball 0 ((1 - sqrt 2 / 2) * r)) \ cball 0 r" + proof - + have "(1 - sqrt 2 / 2) * r \ r" + by (simp add: \0 < r\) + then show ?thesis + by (meson ball_subset_cball closed_cball closure_minimal dual_order.trans subset_ball) + qed + qed have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))" - apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force) - using \0 < r\ mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff) - using False \0 < r\ centre_in_ball holf' holomorphic_nonconstant by blast + proof (rule open_mapping_thm [OF holf' open_ball connected_ball]) + show "interior (ball 0 ((1 - sqrt 2 / 2) * r)) \ ball (0::complex) r" + using \0 < r\ mult_pos_pos sq201 by (simp add: ball_subset_ball_iff) + show "\ f constant_on ball 0 r" + using False \0 < r\ centre_in_ball holf' holomorphic_nonconstant by blast + qed auto have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) = ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))" by simp also have "... \ f ` ball 0 ((1 - sqrt 2 / 2) * r)" proof - have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \ norm (f z)" - if "norm z = (1 - sqrt 2 / 2) * r" for z + if "norm z = (1 - sqrt 2 / 2) * r" for z apply (rule order_trans [OF _ *]) using \0 < r\ - apply (simp_all add: field_simps power2_eq_square that) + apply (simp_all add: field_simps power2_eq_square that) apply (simp add: mult.assoc [symmetric]) done show ?thesis apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball]) - using \0 < r\ sq201 3 apply simp_all - using C_def \0 < C\ sq3 apply force - done - qed + using \0 < r\ sq201 3 C_def \0 < C\ sq3 by auto + qed also have "... \ f ` ball 0 r" - apply (rule image_subsetI [OF imageI], simp) - apply (erule less_le_trans) - using \0 < r\ apply (auto simp: field_simps) - done + proof - + have "\x. (1 - sqrt 2 / 2) * r \ r" + using \0 < r\ by (auto simp: field_simps) + then show ?thesis + by auto + qed finally show ?thesis . qed qed lemma Bloch_lemma: assumes holf: "f holomorphic_on cball a r" and "0 < r" - and le: "\z. z \ ball a r \ norm(deriv f z) \ 2 * norm(deriv f a)" - shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \ f ` ball a r" + and le: "\z. z \ ball a r \ norm(deriv f z) \ 2 * norm(deriv f a)" + shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \ f ` ball a r" (is "?lhs \ ?rhs") proof - have fz: "(\z. f (a + z)) = f o (\z. (a + z))" by (simp add: o_def) have hol0: "(\z. f (a + z)) holomorphic_on cball 0 r" unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+ then have [simp]: "\x. norm x < r \ (\z. f (a + z)) field_differentiable at x" by (metis open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel) have [simp]: "\z. norm z < r \ f field_differentiable at (a + z)" 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) then have [simp]: "f field_differentiable at a" by (metis add.comm_neutral \0 < r\ norm_eq_zero) have hol1: "(\z. f (a + z) - f a) holomorphic_on cball 0 r" by (intro holomorphic_intros hol0) - then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\z. f (a + z) - f a) 0)) - \ (\z. f (a + z) - f a) ` ball 0 r" + then have \
: "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\z. f (a + z) - f a) 0)) + \ (\z. f (a + z) - f a) ` ball 0 r" apply (rule Bloch_lemma_0) - apply (simp_all add: \0 < r\) - apply (simp add: fz deriv_chain) - apply (simp add: dist_norm le) + using \0 < r\ + apply (simp_all add: \0 < r\) + apply (simp add: fz deriv_chain dist_norm le) done - then show ?thesis - apply clarify - apply (drule_tac c="x - f a" in subsetD) - apply (force simp: fz \0 < r\ dist_norm deriv_chain field_differentiable_compose)+ - done + show ?thesis + proof clarify + fix x + assume "x \ ?lhs" + with subsetD [OF \
, of "x - f a"] show "x \ ?rhs" + by (force simp: fz \0 < r\ dist_norm deriv_chain field_differentiable_compose) + qed qed proposition Bloch_unit: assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" obtains b r where "1/12 < r" and "ball b r \ f ` (ball a 1)" proof - define r :: real where "r = 249/256" have "0 < r" "r < 1" by (auto simp: r_def) define g where "g z = deriv f z * of_real(r - norm(z - a))" for z have "deriv f holomorphic_on ball a 1" by (rule holomorphic_deriv [OF holf open_ball]) then have "continuous_on (ball a 1) (deriv f)" using holomorphic_on_imp_continuous_on by blast then have "continuous_on (cball a r) (deriv f)" by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \r < 1\) then have "continuous_on (cball a r) g" by (simp add: g_def continuous_intros) then have 1: "compact (g ` cball a r)" by (rule compact_continuous_image [OF _ compact_cball]) have 2: "g ` cball a r \ {}" using \r > 0\ by auto obtain p where pr: "p \ cball a r" and pge: "\y. y \ cball a r \ norm (g y) \ norm (g p)" using distance_attains_sup [OF 1 2, of 0] by force define t where "t = (r - norm(p - a)) / 2" have "norm (p - a) \ r" using pge [of a] \r > 0\ by (auto simp: g_def norm_mult) then have "norm (p - a) < r" using pr by (simp add: norm_minus_commute dist_norm) then have "0 < t" by (simp add: t_def) have cpt: "cball p t \ ball a r" using \0 < t\ by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \ norm (deriv f p)" if "y \ cball a r" for y proof - have [simp]: "norm (y - a) \ r" using that by (simp add: dist_norm norm_minus_commute) have "norm (g y) \ norm (g p)" using pge [OF that] by simp then have "norm (deriv f y) * abs (r - norm (y - a)) \ norm (deriv f p) * abs (r - norm (p - a))" by (simp only: dist_norm g_def norm_mult norm_of_real) with that \norm (p - a) < r\ show ?thesis by (simp add: dist_norm field_split_simps) qed have le_norm_dfp: "r / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [of a] \r > 0\ by auto have 1: "f holomorphic_on cball p t" - apply (rule holomorphic_on_subset [OF holf]) - using cpt \r < 1\ order_subst1 subset_ball by auto + using cpt \r < 1\ order_subst1 subset_ball + by (force simp add: intro!: holomorphic_on_subset [OF holf]) have 2: "norm (deriv f z) \ 2 * norm (deriv f p)" if "z \ ball p t" for z proof - have z: "z \ cball a r" by (meson ball_subset_cball subsetD cpt that) then have "norm(z - a) < r" by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [OF z] by simp with \norm (z - a) < r\ \norm (p - a) < r\ have "norm (deriv f z) \ (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" - by (simp add: field_simps) + by (simp add: field_simps) also have "... \ 2 * norm (deriv f p)" - apply (rule mult_right_mono) - using that \norm (p - a) < r\ \norm(z - a) < r\ - apply (simp_all add: field_simps t_def dist_norm [symmetric]) - using dist_triangle3 [of z a p] by linarith + proof (rule mult_right_mono) + show "(r - cmod (p - a)) / (r - cmod (z - a)) \ 2" + using that \norm (p - a) < r\ \norm(z - a) < r\ dist_triangle3 [of z a p] + by (simp add: field_simps t_def dist_norm [symmetric]) + qed auto finally show ?thesis . qed have sqrt2: "sqrt 2 < 2113/1494" by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" using sq3 sqrt2 by (auto simp: field_simps r_def) also have "... \ cmod (deriv f p) * (r - cmod (p - a))" using \norm (p - a) < r\ le_norm_dfp by (simp add: pos_divide_le_eq) finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" using pos_divide_less_eq half_gt_zero_iff sq3 by blast then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" using sq3 by (simp add: mult.commute t_def) have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball p t" by (rule Bloch_lemma [OF 1 \0 < t\ 2]) also have "... \ f ` ball a 1" - apply (rule image_mono) - apply (rule order_trans [OF ball_subset_cball]) - apply (rule order_trans [OF cpt]) - using \0 < t\ \r < 1\ apply (simp add: ball_subset_ball_iff dist_norm) - done + proof - + have "ball a r \ ball a 1" + using \0 < t\ \r < 1\ by (simp add: ball_subset_ball_iff dist_norm) + then show ?thesis + using ball_subset_cball cpt by blast + qed finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball a 1" . with ** show ?thesis by (rule that) qed theorem Bloch: assumes holf: "f holomorphic_on ball a r" and "0 < r" and r': "r' \ r * norm (deriv f a) / 12" obtains b where "ball b r' \ f ` (ball a r)" proof (cases "deriv f a = 0") case True with r' show ?thesis using ball_eq_empty that by fastforce next case False define C where "C = deriv f a" have "0 < norm C" using False by (simp add: C_def) have dfa: "f field_differentiable at a" - apply (rule holomorphic_on_imp_differentiable_at [OF holf]) - using \0 < r\ by auto + using \0 < r\ holomorphic_on_imp_differentiable_at [OF holf] by auto have fo: "(\z. f (a + of_real r * z)) = f o (\z. (a + of_real r * z))" by (simp add: o_def) have holf': "f holomorphic_on (\z. a + complex_of_real r * z) ` ball 0 1" - apply (rule holomorphic_on_subset [OF holf]) - using \0 < r\ apply (force simp: dist_norm norm_mult) - done + using \0 < r\ holomorphic_on_subset [OF holf] by (force simp: dist_norm norm_mult) have 1: "(\z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" - apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ - using \0 < r\ by (simp add: C_def False) + using \0 < r\ \0 < norm C\ + by (intro holomorphic_intros holomorphic_on_compose holf'; simp add: fo)+ have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative (deriv f (a + of_real r * z) / C)) (at z)" if "norm z < 1" for z proof - + have fd: "f field_differentiable at (a + complex_of_real r * z)" + using \0 < r\ by (simp_all add: dist_norm norm_mult holomorphic_on_imp_differentiable_at [OF holf] that) have *: "((\x. f (a + of_real r * x)) has_field_derivative (deriv f (a + of_real r * z) * of_real r)) (at z)" - apply (simp add: fo) - apply (rule DERIV_chain [OF field_differentiable_derivI]) - apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) - using \0 < r\ apply (simp add: dist_norm norm_mult that) - apply (rule derivative_eq_intros | simp)+ - done + by (rule fd DERIV_chain [OF field_differentiable_derivI]derivative_eq_intros | simp add: fo)+ show ?thesis apply (rule derivative_eq_intros * | simp)+ using \0 < r\ by (auto simp: C_def False) qed - have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" - apply (subst deriv_cdivide_right) - apply (simp add: field_differentiable_def fo) - apply (rule exI) - apply (rule DERIV_chain [OF field_differentiable_derivI]) - apply (simp add: dfa) - apply (rule derivative_eq_intros | simp add: C_def False fo)+ - using \0 < r\ - apply (simp add: C_def False fo) - apply (simp add: derivative_intros dfa deriv_chain) - done + have "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = deriv (\z. f (a + complex_of_real r * z)) 0 / + (C * complex_of_real r)" + apply (rule deriv_cdivide_right) + by (metis (no_types) DERIV_chain2 add.right_neutral dfa field_differentiable_add_const field_differentiable_def field_differentiable_linear fo mult_zero_right) + also have "... = 1" + using \0 < r\ by (simp add: C_def False fo derivative_intros dfa deriv_chain) + finally have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" . have sb1: "(*) (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 \ f ` ball a r" using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) have sb2: "ball (C * r * b) r' \ (*) (C * r) ` ball b t" if "1 / 12 < t" for b t proof - have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right by auto show ?thesis apply clarify apply (rule_tac x="x / (C * r)" in image_eqI) - using \0 < r\ - apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) - apply (erule less_le_trans) - apply (rule order_trans [OF r' *]) - done + using \0 < r\ apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) + using "*" r' by linarith qed show ?thesis apply (rule Bloch_unit [OF 1 2]) - apply (rename_tac t) apply (rule_tac b="(C * of_real r) * b" in that) - apply (drule image_mono [where f = "\z. (C * of_real r) * z"]) - using sb1 sb2 - apply force - done + using image_mono sb1 sb2 by fastforce qed corollary Bloch_general: - assumes holf: "f holomorphic_on s" and "a \ s" - and tle: "\z. z \ frontier s \ t \ dist a z" + assumes holf: "f holomorphic_on S" and "a \ S" + and tle: "\z. z \ frontier S \ t \ dist a z" and rle: "r \ t * norm(deriv f a) / 12" - obtains b where "ball b r \ f ` s" + obtains b where "ball b r \ f ` S" proof - consider "r \ 0" | "0 < t * norm(deriv f a) / 12" using rle by force then show ?thesis proof cases case 1 then show ?thesis by (simp add: ball_empty that) next case 2 show ?thesis proof (cases "deriv f a = 0") case True then show ?thesis using rle by (simp add: ball_empty that) next case False then have "t > 0" using 2 by (force simp: zero_less_mult_iff) - have "\ ball a t \ s \ ball a t \ frontier s \ {}" - apply (rule connected_Int_frontier [of "ball a t" s], simp_all) - using \0 < t\ \a \ s\ centre_in_ball apply blast - done - with tle have *: "ball a t \ s" by fastforce + have "\ ball a t \ S \ ball a t \ frontier S \ {}" + by (metis Diff_eq_empty_iff \0 < t\ \a \ S\ closure_Int_ball_not_empty closure_subset connected_Int_frontier connected_ball inf.commute) + with tle have *: "ball a t \ S" by fastforce then have 1: "f holomorphic_on ball a t" using holf using holomorphic_on_subset by blast show ?thesis apply (rule Bloch [OF 1 \t > 0\ rle]) apply (rule_tac b=b in that) using * apply force done qed qed qed end