diff --git a/src/HOL/Analysis/Conformal_Mappings.thy b/src/HOL/Analysis/Conformal_Mappings.thy --- a/src/HOL/Analysis/Conformal_Mappings.thy +++ b/src/HOL/Analysis/Conformal_Mappings.thy @@ -1,5096 +1,5086 @@ 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_Theorem begin (* FIXME mv to Cauchy_Integral_Theorem.thy *) subsection\Cauchy's inequality and more versions of Liouville\ lemma Cauchy_higher_deriv_bound: assumes holf: "f holomorphic_on (ball z r)" and contf: "continuous_on (cball z r) f" and fin : "\w. w \ ball z r \ f w \ ball y B0" and "0 < r" and "0 < n" shows "norm ((deriv ^^ n) f z) \ (fact n) * B0 / r^n" proof - have "0 < B0" using \0 < r\ fin [of z] by (metis ball_eq_empty ex_in_conv fin not_less) have le_B0: "\w. cmod (w - z) \ r \ cmod (f w - y) \ B0" apply (rule continuous_on_closure_norm_le [of "ball z r" "\w. f w - y"]) apply (auto simp: \0 < r\ dist_norm norm_minus_commute) apply (rule continuous_intros contf)+ using fin apply (simp add: dist_commute dist_norm less_eq_real_def) done have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w) z - (deriv ^^ n) (\w. y) z" using \0 < n\ by simp also have "... = (deriv ^^ n) (\w. f w - y) z" by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \0 < r\) finally have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w - y) z" . have contf': "continuous_on (cball z r) (\u. f u - y)" by (rule contf continuous_intros)+ have holf': "(\u. (f u - y)) holomorphic_on (ball z r)" by (simp add: holf holomorphic_on_diff) define a where "a = (2 * pi)/(fact n)" have "0 < a" by (simp add: a_def) have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" using \0 < r\ by (simp add: a_def field_split_simps) have der_dif: "(deriv ^^ n) (\w. f w - y) z = (deriv ^^ n) f z" using \0 < r\ \0 < n\ by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) have "norm ((2 * of_real pi * \)/(fact n) * (deriv ^^ n) (\w. f w - y) z) \ (B0/r^(Suc n)) * (2 * pi * r)" apply (rule has_contour_integral_bound_circlepath [of "(\u. (f u - y)/(u - z)^(Suc n))" _ z]) using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] using \0 < B0\ \0 < r\ apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) done then show ?thesis using \0 < r\ by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) qed lemma Cauchy_inequality: assumes holf: "f holomorphic_on (ball \ r)" and contf: "continuous_on (cball \ r) f" and "0 < r" and nof: "\x. norm(\-x) = r \ norm(f x) \ B" shows "norm ((deriv ^^ n) f \) \ (fact n) * B / r^n" proof - obtain x where "norm (\-x) = r" by (metis abs_of_nonneg add_diff_cancel_left' \0 < r\ diff_add_cancel dual_order.strict_implies_order norm_of_real) then have "0 \ B" by (metis nof norm_not_less_zero not_le order_trans) have "((\u. f u / (u - \) ^ Suc n) has_contour_integral (2 * pi) * \ / fact n * (deriv ^^ n) f \) (circlepath \ r)" apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) using \0 < r\ by simp then have "norm ((2 * pi * \)/(fact n) * (deriv ^^ n) f \) \ (B / r^(Suc n)) * (2 * pi * r)" apply (rule has_contour_integral_bound_circlepath) using \0 \ B\ \0 < r\ apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc) done then show ?thesis using \0 < r\ by (simp add: norm_divide norm_mult field_simps) qed lemma Liouville_polynomial: assumes holf: "f holomorphic_on UNIV" and nof: "\z. A \ norm z \ norm(f z) \ B * norm z ^ n" shows "f \ = (\k\n. (deriv^^k) f 0 / fact k * \ ^ k)" proof (cases rule: le_less_linear [THEN disjE]) assume "B \ 0" then have "\z. A \ norm z \ norm(f z) = 0" by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le) then have f0: "(f \ 0) at_infinity" using Lim_at_infinity by force then have [simp]: "f = (\w. 0)" using Liouville_weak [OF holf, of 0] by (simp add: eventually_at_infinity f0) meson show ?thesis by simp next assume "0 < B" have "((\k. (deriv ^^ k) f 0 / (fact k) * (\ - 0)^k) sums f \)" apply (rule holomorphic_power_series [where r = "norm \ + 1"]) using holf holomorphic_on_subset apply auto done then have sumsf: "((\k. (deriv ^^ k) f 0 / (fact k) * \^k) sums f \)" by simp have "(deriv ^^ k) f 0 / fact k * \ ^ k = 0" if "k>n" for k proof (cases "(deriv ^^ k) f 0 = 0") case True then show ?thesis by simp next case False define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" have "1 \ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" using \0 < B\ by simp then have wge1: "1 \ norm w" by (metis norm_of_real w_def) then have "w \ 0" by auto have kB: "0 < fact k * B" using \0 < B\ by simp then have "0 \ fact k * B / cmod ((deriv ^^ k) f 0)" by simp then have wgeA: "A \ cmod w" by (simp only: w_def norm_of_real) have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" using \0 < B\ by simp then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" by (metis norm_of_real w_def) then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" using False by (simp add: field_split_simps mult.commute split: if_split_asm) also have "... \ fact k * (B * norm w ^ n) / norm w ^ k" apply (rule Cauchy_inequality) using holf holomorphic_on_subset apply force using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast using \w \ 0\ apply simp by (metis nof wgeA dist_0_norm dist_norm) also have "... = fact k * (B * 1 / cmod w ^ (k-n))" apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) using \k>n\ \w \ 0\ \0 < B\ apply (simp add: field_split_simps semiring_normalization_rules) done also have "... = fact k * B / cmod w ^ (k-n)" by simp finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . then have "1 / cmod w < 1 / cmod w ^ (k - n)" by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) then have "cmod w ^ (k - n) < cmod w" by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) with self_le_power [OF wge1] have False by (meson diff_is_0_eq not_gr0 not_le that) then show ?thesis by blast qed then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \ ^ (k + Suc n) = 0" for k using not_less_eq by blast then have "(\i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \ ^ (i + Suc n)) sums 0" by (rule sums_0) with sums_split_initial_segment [OF sumsf, where n = "Suc n"] show ?thesis using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce qed text\Every bounded entire function is a constant function.\ theorem Liouville_theorem: assumes holf: "f holomorphic_on UNIV" and bf: "bounded (range f)" obtains c where "\z. f z = c" proof - obtain B where "\z. cmod (f z) \ B" by (meson bf bounded_pos rangeI) then show ?thesis using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast qed text\A holomorphic function f has only isolated zeros unless f is 0.\ lemma powser_0_nonzero: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes r: "0 < r" and sm: "\x. norm (x - \) < r \ (\n. a n * (x - \) ^ n) sums (f x)" and [simp]: "f \ = 0" and m0: "a m \ 0" and "m>0" obtains s where "0 < s" and "\z. z \ cball \ s - {\} \ f z \ 0" proof - have "r \ conv_radius a" using sm sums_summable by (auto simp: le_conv_radius_iff [where \=\]) obtain m where am: "a m \ 0" and az [simp]: "(\n. n a n = 0)" apply (rule_tac m = "LEAST n. a n \ 0" in that) using m0 apply (rule LeastI2) apply (fastforce intro: dest!: not_less_Least)+ done define b where "b i = a (i+m) / a m" for i define g where "g x = suminf (\i. b i * (x - \) ^ i)" for x have [simp]: "b 0 = 1" by (simp add: am b_def) { fix x::'a assume "norm (x - \) < r" then have "(\n. (a m * (x - \)^m) * (b n * (x - \)^n)) sums (f x)" using am az sm sums_zero_iff_shift [of m "(\n. a n * (x - \) ^ n)" "f x"] by (simp add: b_def monoid_mult_class.power_add algebra_simps) then have "x \ \ \ (\n. b n * (x - \)^n) sums (f x / (a m * (x - \)^m))" using am by (simp add: sums_mult_D) } note bsums = this then have "norm (x - \) < r \ summable (\n. b n * (x - \)^n)" for x using sums_summable by (cases "x=\") auto then have "r \ conv_radius b" by (simp add: le_conv_radius_iff [where \=\]) then have "r/2 < conv_radius b" using not_le order_trans r by fastforce then have "continuous_on (cball \ (r/2)) g" using powser_continuous_suminf [of "r/2" b \] by (simp add: g_def) then obtain s where "s>0" "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ dist (g x) (g \) < 1/2" apply (rule continuous_onE [where x=\ and e = "1/2"]) using r apply (auto simp: norm_minus_commute dist_norm) done moreover have "g \ = 1" by (simp add: g_def) ultimately have gnz: "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ (g x) \ 0" by fastforce have "f x \ 0" if "x \ \" "norm (x - \) \ s" "norm (x - \) \ r/2" for x using bsums [of x] that gnz [of x] apply (auto simp: g_def) using r sums_iff by fastforce then show ?thesis apply (rule_tac s="min s (r/2)" in that) using \0 < r\ \0 < s\ by (auto simp: dist_commute dist_norm) qed 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 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) have "0 < min r s" by (simp add: \0 < r\ \0 < s\) 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) by (metis open_ball \\ islimpt T\ centre_in_ball fT0 insertE insert_Diff islimptE) 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 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 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)" 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) 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 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 \))" apply (rule continuous_on_compose2 [OF Complex_Analysis_Basics.continuous_on_norm_id]) using hol frsbU holomorphic_on_imp_continuous_on holomorphic_on_subset by blast+ 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 moreover define \ where "\ \ norm (f w - f \) / 3" ultimately have "0 < \" using \0 < r\ dist_complex_def r sne by auto have "ball (f \) \ \ f ` U" proof fix \ assume \: "\ \ ball (f \) \" have *: "cmod (\ - f \) < cmod (\ - f z)" if "cmod (\ - z) = r" for z proof - have lt: "cmod (f w - f \) / 3 < cmod (\ - f z)" using w [OF that] \ using dist_triangle2 [of "f \" "\" "f z"] dist_triangle2 [of "f \" "f z" \] by (simp add: \_def dist_norm norm_minus_commute) show ?thesis by (metis \_def dist_commute dist_norm less_trans lt mem_ball \) qed have "continuous_on (cball \ r) (\z. \ - f z)" apply (rule continuous_intros)+ using \cball \ r \ U\ \f holomorphic_on U\ apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) done 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 ultimately obtain z where "z \ ball \ r" "\ - f z = 0" apply (rule holomorphic_contract_to_zero) apply (blast intro!: \0 < r\ *)+ done 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 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 } 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 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 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" 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) 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 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) 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 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 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 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) 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) 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 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 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 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 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 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" 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 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 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 "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) 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 \]) 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) 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 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 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 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 - 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 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 from lt1 have "f (inverse x) \ 0 \ x \ 0 \ cmod x < r \ 1 < cmod (f (inverse x))" for x using one_less_inverse by force 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 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 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 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 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 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 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 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 with k m show ?thesis by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right) qed have "((inverse \ f) \ 0) at_infinity" proof (rule Lim_at_infinityI) fix e::real assume "0 < e" with compf [of "cball 0 (inverse e)"] show "\B. \x. B \ cmod x \ dist ((inverse \ f) x) 0 \ e" apply 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 qed then show ?rhs apply (rule pole_at_infinity [OF assms]) using 2 apply blast done 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) 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" 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 \) 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) 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' * 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 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 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 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 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)+ using holg by (simp add: holomorphic_on_subset subset_ball) 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) define U where "U = (\w. (w - \) * g w) ` T" have "open U" by (metis oimT U_def) have "0 \ U" apply (auto simp: U_def) apply (rule image_eqI [where x = \]) apply (auto simp: \\ \ T\) done then 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 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 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')" + 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 - apply (rule has_derivative_inverse_strong [of S x f g ]) - by auto + 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')" + 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 - apply (rule has_derivative_inverse_strong_x [of S g y f]) - by auto + 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) 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 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 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" 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) 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 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 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 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) 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 "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) 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 "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) 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 \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) 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 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 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 (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 "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 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 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 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 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 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 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 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 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 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 apply (rule order_trans [OF _ *]) using \0 < r\ 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 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 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" 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" apply (rule Bloch_lemma_0) apply (simp_all add: \0 < r\) apply (simp add: fz complex_derivative_chain) apply (simp add: 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 complex_derivative_chain field_differentiable_compose)+ done 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 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) 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 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 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 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 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) 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 *: "((\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 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 complex_derivative_chain) done 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 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 qed corollary Bloch_general: 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" 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 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 subsection \Cauchy's residue theorem\ text\Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. Interactive Theorem Proving\ definition\<^marker>\tag important\ residue :: "(complex \ complex) \ complex \ complex" where "residue f z = (SOME int. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" lemma Eps_cong: assumes "\x. P x = Q x" shows "Eps P = Eps Q" using ext[of P Q, OF assms] by simp lemma residue_cong: assumes eq: "eventually (\z. f z = g z) (at z)" and "z = z'" shows "residue f z = residue g z'" proof - from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) let ?P = "\f c e. (\\>0. \ < e \ (f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \))" have "residue f z = residue g z" unfolding residue_def proof (rule Eps_cong) fix c :: complex have "\e>0. ?P g c e" if "\e>0. ?P f c e" and "eventually (\z. f z = g z) (at z)" for f g proof - from that(1) obtain e where e: "e > 0" "?P f c e" by blast from that(2) obtain e' where e': "e' > 0" "\z'. z' \ z \ dist z' z < e' \ f z' = g z'" unfolding eventually_at by blast have "?P g c (min e e')" proof (intro allI exI impI, goal_cases) case (1 \) hence "(f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \)" using e(2) by auto thus ?case proof (rule has_contour_integral_eq) fix z' assume "z' \ path_image (circlepath z \)" hence "dist z' z < e'" and "z' \ z" using 1 by (auto simp: dist_commute) with e'(2)[of z'] show "f z' = g z'" by simp qed qed moreover from e and e' have "min e e' > 0" by auto ultimately show ?thesis by blast qed from this[OF _ eq] and this[OF _ eq'] show "(\e>0. ?P f c e) \ (\e>0. ?P g c e)" by blast qed with assms show ?thesis by simp qed lemma contour_integral_circlepath_eq: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" and e2_cball:"cball z e2 \ s" shows "f contour_integrable_on circlepath z e1" "f contour_integrable_on circlepath z e2" "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" proof - define l where "l \ linepath (z+e2) (z+e1)" have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto have "e2>0" using \e1>0\ \e1\e2\ by auto have zl_img:"z\path_image l" proof assume "z \ path_image l" then have "e2 \ cmod (e2 - e1)" using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \e1>0\ \e2>0\ unfolding l_def by (auto simp add:closed_segment_commute) thus False using \e2>0\ \e1>0\ \e1\e2\ apply (subst (asm) norm_of_real) by auto qed define g where "g \ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" proof - show "f contour_integrable_on circlepath z e2" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e2>0\ e2_cball by auto show "f contour_integrable_on (circlepath z e1)" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e1>0\ \e1\e2\ e2_cball by auto qed have [simp]:"f contour_integrable_on l" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) hence "closed_segment (z + e2) (z + e1) \ s - {z}" using zl_img e2_cball unfolding l_def by auto then show "f contour_integrable_on l" unfolding l_def apply (intro contour_integrable_continuous_linepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) by auto qed let ?ig="\g. contour_integral g f" have "(f has_contour_integral 0) g" proof (rule Cauchy_theorem_global[OF _ f_holo]) show "open (s - {z})" using \open s\ by auto show "valid_path g" unfolding g_def l_def by auto show "pathfinish g = pathstart g" unfolding g_def l_def by auto next have path_img:"path_image g \ cball z e2" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) moreover have "sphere z \e1\ \ cball z e2" using \e2>0\ \e1\e2\ \e1>0\ by auto ultimately show ?thesis unfolding g_def l_def using \e2>0\ by (simp add: path_image_join closed_segment_commute) qed show "path_image g \ s - {z}" proof - have "z\path_image g" using zl_img unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) moreover note \cball z e2 \ s\ and path_img ultimately show ?thesis by auto qed show "winding_number g w = 0" when"w \ s - {z}" for w proof - have "winding_number g w = 0" when "w\s" using that e2_cball apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) by (auto simp add:g_def l_def) moreover have "winding_number g z=0" proof - let ?Wz="\g. winding_number g z" have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + ?Wz (reversepath l)" using \e2>0\ \e1>0\ zl_img unfolding g_def l_def by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" using zl_img apply (subst (2) winding_number_reversepath) by (auto simp add:l_def closed_segment_commute) also have "... = 0" proof - have "?Wz (circlepath z e2) = 1" using \e2>0\ by (auto intro: winding_number_circlepath_centre) moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \e1>0\ apply (subst winding_number_reversepath) by (auto intro: winding_number_circlepath_centre) ultimately show ?thesis by auto qed finally show ?thesis . qed ultimately show ?thesis using that by auto qed qed then have "0 = ?ig g" using contour_integral_unique by simp also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + ?ig (reversepath l)" unfolding g_def by (auto simp add:contour_integrable_reversepath_eq) also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" by (auto simp add:contour_integral_reversepath) finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" by simp qed lemma base_residue: assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" and r_cball:"cball z r \ s" shows "(f has_contour_integral 2 * pi * \ * (residue f z)) (circlepath z r)" proof - obtain e where "e>0" and e_cball:"cball z e \ s" using open_contains_cball[of s] \open s\ \z\s\ by auto define c where "c \ 2 * pi * \" define i where "i \ contour_integral (circlepath z e) f / c" have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ proof - have "contour_integral (circlepath z e) f = contour_integral (circlepath z \) f" "f contour_integrable_on circlepath z \" "f contour_integrable_on circlepath z e" using \\ by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ then show ?thesis unfolding i_def c_def by (auto intro:has_contour_integral_integral) qed then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" unfolding residue_def c_def apply (rule_tac someI[of _ i],intro exI[where x=e]) by (auto simp add:\e>0\ c_def) then obtain e' where "e'>0" and e'_def:"\\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" by auto let ?int="\e. contour_integral (circlepath z e) f" define \ where "\ \ Min {r,e'} / 2" have "\>0" "\\r" "\r>0\ \e'>0\ unfolding \_def by auto have "(f has_contour_integral c * (residue f z)) (circlepath z \)" using e'_def[rule_format,OF \\>0\ \\] . then show ?thesis unfolding c_def using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \" "circlepath z r"]) qed lemma residue_holo: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s" shows "residue f z = 0" proof - define c where "c \ 2 * pi * \" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f has_contour_integral c*residue f z) (circlepath z e)" using f_holo by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) moreover have "(f has_contour_integral 0) (circlepath z e)" using f_holo e_cball \e>0\ by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) ultimately have "c*residue f z =0" using has_contour_integral_unique by blast thus ?thesis unfolding c_def by auto qed lemma residue_const:"residue (\_. c) z = 0" by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) lemma residue_add: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z + g z) z= residue f z + residue g z" proof - define c where "c \ 2 * pi * \" define fg where "fg \ (\z. f z+g z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(fg has_contour_integral c * residue fg z) (circlepath z e)" unfolding fg_def using f_holo g_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro:holomorphic_intros) moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" unfolding fg_def using f_holo g_holo by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) ultimately have "c*(residue f z + residue g z) = c * residue fg z" using has_contour_integral_unique by (auto simp add:distrib_left) thus ?thesis unfolding fg_def by (auto simp add:c_def) qed lemma residue_lmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. c * (f z)) z= c * residue f z" proof (cases "c=0") case True thus ?thesis using residue_const by auto next case False define c' where "c' \ 2 * pi * \" define f' where "f' \ (\z. c * (f z))" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) by (auto intro:holomorphic_intros) moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" unfolding f'_def using f_holo by (auto intro: has_contour_integral_lmul base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) ultimately have "c' * residue f' z = c * (c' * residue f z)" using has_contour_integral_unique by auto thus ?thesis unfolding f'_def c'_def using False by (auto simp add:field_simps) qed lemma residue_rmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) * c) z= residue f z * c" using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) lemma residue_div: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) / c) z= residue f z / c " using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) lemma residue_neg: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. - (f z)) z= - residue f z" using residue_lmul[OF assms,of "-1"] by auto lemma residue_diff: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z - g z) z= residue f z - residue g z" using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] by (auto intro:holomorphic_intros g_holo) lemma residue_simple: assumes "open s" "z\s" and f_holo:"f holomorphic_on s" shows "residue (\w. f w / (w - z)) z = f z" proof - define c where "c \ 2 * pi * \" define f' where "f' \ \w. f w / (w - z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c * f z) (circlepath z e)" unfolding f'_def c_def using \e>0\ f_holo e_cball by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro!:holomorphic_intros) ultimately have "c * f z = c * residue f' z" using has_contour_integral_unique by blast thus ?thesis unfolding c_def f'_def by auto qed lemma residue_simple': assumes s: "open s" "z \ s" and holo: "f holomorphic_on (s - {z})" and lim: "((\w. f w * (w - z)) \ c) (at z)" shows "residue f z = c" proof - define g where "g = (\w. if w = z then c else f w * (w - z))" from holo have "(\w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P") by (force intro: holomorphic_intros) also have "?P \ g holomorphic_on (s - {z})" by (intro holomorphic_cong refl) (simp_all add: g_def) finally have *: "g holomorphic_on (s - {z})" . note lim also have "(\w. f w * (w - z)) \z\ c \ g \z\ g z" by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter) finally have **: "g \z\ g z" . have g_holo: "g holomorphic_on s" by (rule no_isolated_singularity'[where K = "{z}"]) (insert assms * **, simp_all add: at_within_open_NO_MATCH) from s and this have "residue (\w. g w / (w - z)) z = g z" by (rule residue_simple) also have "\\<^sub>F za in at z. g za / (za - z) = f za" unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def) hence "residue (\w. g w / (w - z)) z = residue f z" by (intro residue_cong refl) finally show ?thesis by (simp add: g_def) qed lemma residue_holomorphic_over_power: assumes "open A" "z0 \ A" "f holomorphic_on A" shows "residue (\z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n" proof - let ?f = "\z. f z / (z - z0) ^ Suc n" from assms(1,2) obtain r where r: "r > 0" "cball z0 r \ A" by (auto simp: open_contains_cball) have "(?f has_contour_integral 2 * pi * \ * residue ?f z0) (circlepath z0 r)" using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros) moreover have "(?f has_contour_integral 2 * pi * \ / fact n * (deriv ^^ n) f z0) (circlepath z0 r)" using assms r by (intro Cauchy_has_contour_integral_higher_derivative_circlepath) (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on) ultimately have "2 * pi * \ * residue ?f z0 = 2 * pi * \ / fact n * (deriv ^^ n) f z0" by (rule has_contour_integral_unique) thus ?thesis by (simp add: field_simps) qed lemma residue_holomorphic_over_power': assumes "open A" "0 \ A" "f holomorphic_on A" shows "residue (\z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" using residue_holomorphic_over_power[OF assms] by simp lemma get_integrable_path: assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\s-pts" "b\s-pts" obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" "path_image g \ s-pts" "f contour_integrable_on g" using assms proof (induct arbitrary:s thesis a rule:finite_induct[OF \finite pts\]) case 1 obtain g where "valid_path g" "path_image g \ s" "pathstart g = a" "pathfinish g = b" using connected_open_polynomial_connected[OF \open s\,of a b ] \connected (s - {})\ valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto moreover have "f contour_integrable_on g" using contour_integrable_holomorphic_simple[OF _ \open s\ \valid_path g\ \path_image g \ s\,of f] \f holomorphic_on s - {}\ by auto ultimately show ?case using "1"(1)[of g] by auto next case idt:(2 p pts) obtain e where "e>0" and e:"\w\ball a e. w \ s \ (w \ a \ w \ insert p pts)" using finite_ball_avoid[OF \open s\ \finite (insert p pts)\, of a] \a \ s - insert p pts\ by auto define a' where "a' \ a+e/2" have "a'\s-{p} -pts" using e[rule_format,of "a+e/2"] \e>0\ by (auto simp add:dist_complex_def a'_def) then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" "path_image g' \ s - {p} - pts" "f contour_integrable_on g'" using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) by (metis Diff_insert2 open_delete) define g where "g \ linepath a a' +++ g'" have "valid_path g" unfolding g_def by (auto intro: valid_path_join) moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto moreover have "path_image g \ s - insert p pts" unfolding g_def proof (rule subset_path_image_join) have "closed_segment a a' \ ball a e" using \e>0\ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) then show "path_image (linepath a a') \ s - insert p pts" using e idt(9) by auto next show "path_image g' \ s - insert p pts" using g'(4) by blast qed moreover have "f contour_integrable_on g" proof - have "closed_segment a a' \ ball a e" using \e>0\ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) then have "continuous_on (closed_segment a a') f" using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] apply (elim continuous_on_subset) by auto then have "f contour_integrable_on linepath a a'" using contour_integrable_continuous_linepath by auto then show ?thesis unfolding g_def apply (rule contour_integrable_joinI) by (auto simp add: \e>0\) qed ultimately show ?case using idt.prems(1)[of g] by auto qed lemma Cauchy_theorem_aux: assumes "open s" "connected (s-pts)" "finite pts" "pts \ s" "f holomorphic_on s-pts" "valid_path g" "pathfinish g = pathstart g" "path_image g \ s-pts" "\z. (z \ s) \ winding_number g z = 0" "\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" using assms proof (induct arbitrary:s g rule:finite_induct[OF \finite pts\]) case 1 then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique) next case (2 p pts) note fin[simp] = \finite (insert p pts)\ and connected = \connected (s - insert p pts)\ and valid[simp] = \valid_path g\ and g_loop[simp] = \pathfinish g = pathstart g\ and holo[simp]= \f holomorphic_on s - insert p pts\ and path_img = \path_image g \ s - insert p pts\ and winding = \\z. z \ s \ winding_number g z = 0\ and h = \\pa\s. 0 < h pa \ (\w\cball pa (h pa). w \ s \ (w \ pa \ w \ insert p pts))\ have "h p>0" and "p\s" and h_p: "\w\cball p (h p). w \ s \ (w \ p \ w \ insert p pts)" using h \insert p pts \ s\ by auto obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" "path_image pg \ s-insert p pts" "f contour_integrable_on pg" proof - have "p + h p\cball p (h p)" using h[rule_format,of p] by (simp add: \p \ s\ dist_norm) then have "p + h p \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ by fastforce moreover have "pathstart g \ s - insert p pts " using path_img by auto ultimately show ?thesis using get_integrable_path[OF \open s\ connected fin holo,of "pathstart g" "p+h p"] that by blast qed obtain n::int where "n=winding_number g p" using integer_winding_number[OF _ g_loop,of p] valid path_img by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) define p_circ where "p_circ \ circlepath p (h p)" define p_circ_pt where "p_circ_pt \ linepath (p+h p) (p+h p)" define n_circ where "n_circ \ \n. ((+++) p_circ ^^ n) p_circ_pt" define cp where "cp \ if n\0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" have n_circ:"valid_path (n_circ k)" "winding_number (n_circ k) p = k" "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" "path_image (n_circ k) = (if k=0 then {p + h p} else sphere p (h p))" "p \ path_image (n_circ k)" "\p'. p'\s - pts \ winding_number (n_circ k) p'=0 \ p'\path_image (n_circ k)" "f contour_integrable_on (n_circ k)" "contour_integral (n_circ k) f = k * contour_integral p_circ f" for k proof (induct k) case 0 show "valid_path (n_circ 0)" and "path_image (n_circ 0) = (if 0=0 then {p + h p} else sphere p (h p))" and "winding_number (n_circ 0) p = of_nat 0" and "pathstart (n_circ 0) = p + h p" and "pathfinish (n_circ 0) = p + h p" and "p \ path_image (n_circ 0)" unfolding n_circ_def p_circ_pt_def using \h p > 0\ by (auto simp add: dist_norm) show "winding_number (n_circ 0) p'=0 \ p'\path_image (n_circ 0)" when "p'\s- pts" for p' unfolding n_circ_def p_circ_pt_def apply (auto intro!:winding_number_trivial) by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+ show "f contour_integrable_on (n_circ 0)" unfolding n_circ_def p_circ_pt_def by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing) show "contour_integral (n_circ 0) f = of_nat 0 * contour_integral p_circ f" unfolding n_circ_def p_circ_pt_def by auto next case (Suc k) have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto have pcirc:"p \ path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" using Suc(3) unfolding p_circ_def using \h p > 0\ by (auto simp add: p_circ_def) have pcirc_image:"path_image p_circ \ s - insert p pts" proof - have "path_image p_circ \ cball p (h p)" using \0 < h p\ p_circ_def by auto then show ?thesis using h_p pcirc(1) by auto qed have pcirc_integrable:"f contour_integrable_on p_circ" by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on holomorphic_on_subset[OF holo]) show "valid_path (n_circ (Suc k))" using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto show "path_image (n_circ (Suc k)) = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" proof - have "path_image p_circ = sphere p (h p)" unfolding p_circ_def using \0 < h p\ by auto then show ?thesis unfolding n_Suc using Suc.hyps(5) \h p>0\ by (auto simp add: path_image_join[OF pcirc(3)] dist_norm) qed then show "p \ path_image (n_circ (Suc k))" using \h p>0\ by auto show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" proof - have "winding_number p_circ p = 1" by (simp add: \h p > 0\ p_circ_def winding_number_circlepath_centre) moreover have "p \ path_image (n_circ k)" using Suc(5) \h p>0\ by auto then have "winding_number (p_circ +++ n_circ k) p = winding_number p_circ p + winding_number (n_circ k) p" using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc apply (intro winding_number_join) by auto ultimately show ?thesis using Suc(2) unfolding n_circ_def by auto qed show "pathstart (n_circ (Suc k)) = p + h p" by (simp add: n_circ_def p_circ_def) show "pathfinish (n_circ (Suc k)) = p + h p" using Suc(4) unfolding n_circ_def by auto show "winding_number (n_circ (Suc k)) p'=0 \ p'\path_image (n_circ (Suc k))" when "p'\s-pts" for p' proof - have " p' \ path_image p_circ" using \p \ s\ h p_circ_def that using pcirc_image by blast moreover have "p' \ path_image (n_circ k)" using Suc.hyps(7) that by blast moreover have "winding_number p_circ p' = 0" proof - have "path_image p_circ \ cball p (h p)" using h unfolding p_circ_def using \p \ s\ by fastforce moreover have "p'\cball p (h p)" using \p \ s\ h that "2.hyps"(2) by fastforce ultimately show ?thesis unfolding p_circ_def apply (intro winding_number_zero_outside) by auto qed ultimately show ?thesis unfolding n_Suc apply (subst winding_number_join) by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that]) qed show "f contour_integrable_on (n_circ (Suc k))" unfolding n_Suc by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]) show "contour_integral (n_circ (Suc k)) f = (Suc k) * contour_integral p_circ f" unfolding n_Suc by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] Suc(9) algebra_simps) qed have cp[simp]:"pathstart cp = p + h p" "pathfinish cp = p + h p" "valid_path cp" "path_image cp \ s - insert p pts" "winding_number cp p = - n" "\p'. p'\s - pts \ winding_number cp p'=0 \ p' \ path_image cp" "f contour_integrable_on cp" "contour_integral cp f = - n * contour_integral p_circ f" proof - show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp" using n_circ unfolding cp_def by auto next have "sphere p (h p) \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ by force moreover have "p + complex_of_real (h p) \ s - insert p pts" using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE) ultimately show "path_image cp \ s - insert p pts" unfolding cp_def using n_circ(5) by auto next show "winding_number cp p = - n" unfolding cp_def using winding_number_reversepath n_circ \h p>0\ by (auto simp: valid_path_imp_path) next show "winding_number cp p'=0 \ p' \ path_image cp" when "p'\s - pts" for p' unfolding cp_def apply (auto) apply (subst winding_number_reversepath) by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1)) next show "f contour_integrable_on cp" unfolding cp_def using contour_integrable_reversepath_eq n_circ(1,8) by auto next show "contour_integral cp f = - n * contour_integral p_circ f" unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9) by auto qed define g' where "g' \ g +++ pg +++ cp +++ (reversepath pg)" have "contour_integral g' f = (\p\pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \finite pts\ ]) show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) show "open (s - {p})" using \open s\ by auto show " pts \ s - {p}" using \insert p pts \ s\ \ p \ pts\ by blast show "f holomorphic_on s - {p} - pts" using holo \p \ pts\ by (metis Diff_insert2) show "valid_path g'" unfolding g'_def cp_def using n_circ valid pg g_loop by (auto intro!:valid_path_join ) show "pathfinish g' = pathstart g'" unfolding g'_def cp_def using pg(2) by simp show "path_image g' \ s - {p} - pts" proof - define s' where "s' \ s - {p} - pts" have s':"s' = s-insert p pts " unfolding s'_def by auto then show ?thesis using path_img pg(4) cp(4) unfolding g'_def apply (fold s'_def s') apply (intro subset_path_image_join) by auto qed note path_join_imp[simp] show "\z. z \ s - {p} \ winding_number g' z = 0" proof clarify fix z assume z:"z\s - {p}" have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z + winding_number (pg +++ cp +++ (reversepath pg)) z" proof (rule winding_number_join) show "path g" using \valid_path g\ by (simp add: valid_path_imp_path) show "z \ path_image g" using z path_img by auto show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp by (simp add: valid_path_imp_path) next have "path_image (pg +++ cp +++ reversepath pg) \ s - insert p pts" using pg(4) cp(4) by (auto simp:subset_path_image_join) then show "z \ path_image (pg +++ cp +++ reversepath pg)" using z by auto next show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto qed also have "... = winding_number g z + (winding_number pg z + winding_number (cp +++ (reversepath pg)) z)" proof (subst add_left_cancel,rule winding_number_join) show "path pg" and "path (cp +++ reversepath pg)" and "pathfinish pg = pathstart (cp +++ reversepath pg)" by (auto simp add: valid_path_imp_path) show "z \ path_image pg" using pg(4) z by blast show "z \ path_image (cp +++ reversepath pg)" using z by (metis Diff_iff \z \ path_image pg\ contra_subsetD cp(4) insertI1 not_in_path_image_join path_image_reversepath singletonD) qed also have "... = winding_number g z + (winding_number pg z + (winding_number cp z + winding_number (reversepath pg) z))" apply (auto intro!:winding_number_join simp: valid_path_imp_path) apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z) by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) also have "... = winding_number g z + winding_number cp z" apply (subst winding_number_reversepath) apply (auto simp: valid_path_imp_path) by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z) finally have "winding_number g' z = winding_number g z + winding_number cp z" unfolding g'_def . moreover have "winding_number g z + winding_number cp z = 0" using winding z \n=winding_number g p\ by auto ultimately show "winding_number g' z = 0" unfolding g'_def by auto qed show "\pa\s - {p}. 0 < h pa \ (\w\cball pa (h pa). w \ s - {p} \ (w \ pa \ w \ pts))" using h by fastforce qed moreover have "contour_integral g' f = contour_integral g f - winding_number g p * contour_integral p_circ f" proof - have "contour_integral g' f = contour_integral g f + contour_integral (pg +++ cp +++ reversepath pg) f" unfolding g'_def apply (subst contour_integral_join) by (auto simp add:open_Diff[OF \open s\,OF finite_imp_closed[OF fin]] intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral pg f + contour_integral (cp +++ reversepath pg) f" apply (subst contour_integral_join) by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral pg f + contour_integral cp f + contour_integral (reversepath pg) f" apply (subst contour_integral_join) by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral cp f" using contour_integral_reversepath by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f" using \n=winding_number g p\ by auto finally show ?thesis . qed moreover have "winding_number g' p' = winding_number g p'" when "p'\pts" for p' proof - have [simp]: "p' \ path_image g" "p' \ path_image pg" "p'\path_image cp" using "2.prems"(8) that apply blast apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that) by (meson DiffD2 cp(4) rev_subsetD subset_insertI that) have "winding_number g' p' = winding_number g p' + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def apply (subst winding_number_join) apply (simp_all add: valid_path_imp_path) apply (intro not_in_path_image_join) by auto also have "... = winding_number g p' + winding_number pg p' + winding_number (cp +++ reversepath pg) p'" apply (subst winding_number_join) apply (simp_all add: valid_path_imp_path) apply (intro not_in_path_image_join) by auto also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p' + winding_number (reversepath pg) p'" apply (subst winding_number_join) by (simp_all add: valid_path_imp_path) also have "... = winding_number g p' + winding_number cp p'" apply (subst winding_number_reversepath) by (simp_all add: valid_path_imp_path) also have "... = winding_number g p'" using that by auto finally show ?thesis . qed ultimately show ?case unfolding p_circ_def apply (subst (asm) sum.cong[OF refl, of pts _ "\p. winding_number g p * contour_integral (circlepath p (h p)) f"]) by (auto simp add:sum.insert[OF \finite pts\ \p\pts\] algebra_simps) qed lemma Cauchy_theorem_singularities: assumes "open s" "connected s" "finite pts" and holo:"f holomorphic_on s-pts" and "valid_path g" and loop:"pathfinish g = pathstart g" and "path_image g \ s-pts" and homo:"\z. (z \ s) \ winding_number g z = 0" and avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" (is "?L=?R") proof - define circ where "circ \ \p. winding_number g p * contour_integral (circlepath p (h p)) f" define pts1 where "pts1 \ pts \ s" define pts2 where "pts2 \ pts - pts1" have "pts=pts1 \ pts2" "pts1 \ pts2 = {}" "pts2 \ s={}" "pts1\s" unfolding pts1_def pts2_def by auto have "contour_integral g f = (\p\pts1. circ p)" unfolding circ_def proof (rule Cauchy_theorem_aux[OF \open s\ _ _ \pts1\s\ _ \valid_path g\ loop _ homo]) have "finite pts1" unfolding pts1_def using \finite pts\ by auto then show "connected (s - pts1)" using \open s\ \connected s\ connected_open_delete_finite[of s] by auto next show "finite pts1" using \pts = pts1 \ pts2\ assms(3) by auto show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def) show "path_image g \ s - pts1" using assms(7) pts1_def by auto show "\p\s. 0 < h p \ (\w\cball p (h p). w \ s \ (w \ p \ w \ pts1))" by (simp add: avoid pts1_def) qed moreover have "sum circ pts2=0" proof - have "winding_number g p=0" when "p\pts2" for p using \pts2 \ s={}\ that homo[rule_format,of p] by auto thus ?thesis unfolding circ_def apply (intro sum.neutral) by auto qed moreover have "?R=sum circ pts1 + sum circ pts2" unfolding circ_def using sum.union_disjoint[OF _ _ \pts1 \ pts2 = {}\] \finite pts\ \pts=pts1 \ pts2\ by blast ultimately show ?thesis apply (fold circ_def) by auto qed theorem Residue_theorem: fixes s pts::"complex set" and f::"complex \ complex" and g::"real \ complex" assumes "open s" "connected s" "finite pts" and holo:"f holomorphic_on s-pts" and "valid_path g" and loop:"pathfinish g = pathstart g" and "path_image g \ s-pts" and homo:"\z. (z \ s) \ winding_number g z = 0" shows "contour_integral g f = 2 * pi * \ *(\p\pts. winding_number g p * residue f p)" proof - define c where "c \ 2 * pi * \" obtain h where avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" using finite_cball_avoid[OF \open s\ \finite pts\] by metis have "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" using Cauchy_theorem_singularities[OF assms avoid] . also have "... = (\p\pts. c * winding_number g p * residue f p)" proof (intro sum.cong) show "pts = pts" by simp next fix x assume "x \ pts" show "winding_number g x * contour_integral (circlepath x (h x)) f = c * winding_number g x * residue f x" proof (cases "x\s") case False then have "winding_number g x=0" using homo by auto thus ?thesis by auto next case True have "contour_integral (circlepath x (h x)) f = c* residue f x" using \x\pts\ \finite pts\ avoid[rule_format,OF True] apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def]) by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \open s\ finite_imp_closed]) then show ?thesis by auto qed qed also have "... = c * (\p\pts. winding_number g p * residue f p)" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis unfolding c_def . qed subsection \Non-essential singular points\ definition\<^marker>\tag important\ is_pole :: "('a::topological_space \ 'b::real_normed_vector) \ 'a \ bool" where "is_pole f a = (LIM x (at a). f x :> at_infinity)" lemma is_pole_cong: assumes "eventually (\x. f x = g x) (at a)" "a=b" shows "is_pole f a \ is_pole g b" unfolding is_pole_def using assms by (intro filterlim_cong,auto) lemma is_pole_transform: assumes "is_pole f a" "eventually (\x. f x = g x) (at a)" "a=b" shows "is_pole g b" using is_pole_cong assms by auto lemma is_pole_tendsto: fixes f::"('a::topological_space \ 'b::real_normed_div_algebra)" shows "is_pole f x \ ((inverse o f) \ 0) (at x)" unfolding is_pole_def by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at) lemma is_pole_inverse_holomorphic: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and pole:"is_pole f z" and non_z:"\x\s-{z}. f x\0" shows "(\x. if x=z then 0 else inverse (f x)) holomorphic_on s" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" have "isCont g z" unfolding isCont_def using is_pole_tendsto[OF pole] apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \ f"]) by (simp_all add:g_def) moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def by (auto elim!:continuous_on_inverse simp add:non_z) hence "continuous_on (s-{z}) g" unfolding g_def apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"]) by auto ultimately have "continuous_on s g" using open_delete[OF \open s\] \open s\ by (auto simp add:continuous_on_eq_continuous_at) moreover have "(inverse o f) holomorphic_on (s-{z})" unfolding comp_def using f_holo by (auto elim!:holomorphic_on_inverse simp add:non_z) hence "g holomorphic_on (s-{z})" apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"]) by (auto simp add:g_def) ultimately show ?thesis unfolding g_def using \open s\ by (auto elim!: no_isolated_singularity) qed lemma not_is_pole_holomorphic: assumes "open A" "x \ A" "f holomorphic_on A" shows "\is_pole f x" proof - have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at) hence "f \x\ f x" by (simp add: isCont_def) thus "\is_pole f x" unfolding is_pole_def using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto qed lemma is_pole_inverse_power: "n > 0 \ is_pole (\z::complex. 1 / (z - a) ^ n) a" unfolding is_pole_def inverse_eq_divide [symmetric] by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros) (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros) lemma is_pole_inverse: "is_pole (\z::complex. 1 / (z - a)) a" using is_pole_inverse_power[of 1 a] by simp lemma is_pole_divide: fixes f :: "'a :: t2_space \ 'b :: real_normed_field" assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \ 0" shows "is_pole (\z. f z / g z) z" proof - have "filterlim (\z. f z * inverse (g z)) at_infinity (at z)" by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"] filterlim_compose[OF filterlim_inverse_at_infinity])+ (insert assms, auto simp: isCont_def) thus ?thesis by (simp add: field_split_simps is_pole_def) qed lemma is_pole_basic: assumes "f holomorphic_on A" "open A" "z \ A" "f z \ 0" "n > 0" shows "is_pole (\w. f w / (w - z) ^ n) z" proof (rule is_pole_divide) have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at) have "filterlim (\w. (w - z) ^ n) (nhds 0) (at z)" using assms by (auto intro!: tendsto_eq_intros) thus "filterlim (\w. (w - z) ^ n) (at 0) (at z)" by (intro filterlim_atI tendsto_eq_intros) (insert assms, auto simp: eventually_at_filter) qed fact+ lemma is_pole_basic': assumes "f holomorphic_on A" "open A" "0 \ A" "f 0 \ 0" "n > 0" shows "is_pole (\w. f w / w ^ n) 0" using is_pole_basic[of f A 0] assms by simp text \The proposition \<^term>\\x. ((f::complex\complex) \ x) (at z) \ is_pole f z\ can be interpreted as the complex function \<^term>\f\ has a non-essential singularity at \<^term>\z\ (i.e. the singularity is either removable or a pole).\ definition not_essential::"[complex \ complex, complex] \ bool" where "not_essential f z = (\x. f\z\x \ is_pole f z)" definition isolated_singularity_at::"[complex \ complex, complex] \ bool" where "isolated_singularity_at f z = (\r>0. f analytic_on ball z r-{z})" named_theorems singularity_intros "introduction rules for singularities" lemma holomorphic_factor_unique: fixes f::"complex \ complex" and z::complex and r::real and m n::int assumes "r>0" "g z\0" "h z\0" and asm:"\w\ball z r-{z}. f w = g w * (w-z) powr n \ g w\0 \ f w = h w * (w - z) powr m \ h w\0" and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" shows "n=m" proof - have [simp]:"at z within ball z r \ bot" using \r>0\ by (auto simp add:at_within_ball_bot_iff) have False when "n>m" proof - have "(h \ 0) (at z within ball z r)" proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (n - m) * g w"]) have "\w\ball z r-{z}. h w = (w-z)powr(n-m) * g w" using \n>m\ asm \r>0\ apply (auto simp add:field_simps powr_diff) by force then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) powr (n - m) * g x' = h x'" for x' by auto next define F where "F \ at z within ball z r" define f' where "f' \ \x. (x - z) powr (n-m)" have "f' z=0" using \n>m\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def continuous_def apply (subst Lim_ident_at) using \n>m\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) ultimately have "(f' \ 0) F" unfolding F_def by (simp add: continuous_within) moreover have "(g \ g z) F" using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \r>0\ unfolding F_def by auto ultimately show " ((\w. f' w * g w) \ 0) F" using tendsto_mult by fastforce qed moreover have "(h \ h z) (at z within ball z r)" using holomorphic_on_imp_continuous_on[OF h_holo] by (auto simp add:continuous_on_def \r>0\) ultimately have "h z=0" by (auto intro!: tendsto_unique) thus False using \h z\0\ by auto qed moreover have False when "m>n" proof - have "(g \ 0) (at z within ball z r)" proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (m - n) * h w"]) have "\w\ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \m>n\ asm apply (auto simp add:field_simps powr_diff) by force then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) powr (m - n) * h x' = g x'" for x' by auto next define F where "F \ at z within ball z r" define f' where "f' \\x. (x - z) powr (m-n)" have "f' z=0" using \m>n\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def continuous_def apply (subst Lim_ident_at) using \m>n\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) ultimately have "(f' \ 0) F" unfolding F_def by (simp add: continuous_within) moreover have "(h \ h z) F" using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \r>0\ unfolding F_def by auto ultimately show " ((\w. f' w * h w) \ 0) F" using tendsto_mult by fastforce qed moreover have "(g \ g z) (at z within ball z r)" using holomorphic_on_imp_continuous_on[OF g_holo] by (auto simp add:continuous_on_def \r>0\) ultimately have "g z=0" by (auto intro!: tendsto_unique) thus False using \g z\0\ by auto qed ultimately show "n=m" by fastforce qed lemma holomorphic_factor_puncture: assumes f_iso:"isolated_singularity_at f z" and "not_essential f z" \ \\<^term>\f\ has either a removable singularity or a pole at \<^term>\z\\ and non_zero:"\\<^sub>Fw in (at z). f w\0" \ \\<^term>\f\ will not be constantly zero in a neighbour of \<^term>\z\\ shows "\!n::int. \g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r-{z}. f w = g w * (w-z) powr n \ g w\0)" proof - define P where "P = (\f n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have imp_unique:"\!n::int. \g r. P f n g r" when "\n g r. P f n g r" proof (rule ex_ex1I[OF that]) fix n1 n2 :: int assume g1_asm:"\g1 r1. P f n1 g1 r1" and g2_asm:"\g2 r2. P f n2 g2 r2" define fac where "fac \ \n g r. \w\cball z r-{z}. f w = g w * (w - z) powr (of_int n) \ g w \ 0" obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\0" and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\0" and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def by auto define r where "r \ min r1 r2" have "r>0" using \r1>0\ \r2>0\ unfolding r_def by auto moreover have "\w\ball z r-{z}. f w = g1 w * (w-z) powr n1 \ g1 w\0 \ f w = g2 w * (w - z) powr n2 \ g2 w\0" using \fac n1 g1 r1\ \fac n2 g2 r2\ unfolding fac_def r_def by fastforce ultimately show "n1=n2" using g1_holo g2_holo \g1 z\0\ \g2 z\0\ apply (elim holomorphic_factor_unique) by (auto simp add:r_def) qed have P_exist:"\ n g r. P h n g r" when "\z'. (h \ z') (at z)" "isolated_singularity_at h z" "\\<^sub>Fw in (at z). h w\0" for h proof - from that(2) obtain r where "r>0" "h analytic_on ball z r - {z}" unfolding isolated_singularity_at_def by auto obtain z' where "(h \ z') (at z)" using \\z'. (h \ z') (at z)\ by auto define h' where "h'=(\x. if x=z then z' else h x)" have "h' holomorphic_on ball z r" apply (rule no_isolated_singularity'[of "{z}"]) subgoal by (metis LIM_equal Lim_at_imp_Lim_at_within \h \z\ z'\ empty_iff h'_def insert_iff) subgoal using \h analytic_on ball z r - {z}\ analytic_imp_holomorphic h'_def holomorphic_transform by fastforce by auto have ?thesis when "z'=0" proof - have "h' z=0" using that unfolding h'_def by auto moreover have "\ h' constant_on ball z r" using \\\<^sub>Fw in (at z). h w\0\ unfolding constant_on_def frequently_def eventually_at h'_def apply simp by (metis \0 < r\ centre_in_ball dist_commute mem_ball that) moreover note \h' holomorphic_on ball z r\ ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \ ball z r" and g:"g holomorphic_on ball z r1" "\w. w \ ball z r1 \ h' w = (w - z) ^ n * g w" "\w. w \ ball z r1 \ g w \ 0" using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified, OF \h' holomorphic_on ball z r\ \r>0\ \h' z=0\ \\ h' constant_on ball z r\] by (auto simp add:dist_commute) define rr where "rr=r1/2" have "P h' n g rr" unfolding P_def rr_def using \n>0\ \r1>0\ g by (auto simp add:powr_nat) then have "P h n g rr" unfolding h'_def P_def by auto then show ?thesis unfolding P_def by blast qed moreover have ?thesis when "z'\0" proof - have "h' z\0" using that unfolding h'_def by auto obtain r1 where "r1>0" "cball z r1 \ ball z r" "\x\cball z r1. h' x\0" proof - have "isCont h' z" "h' z\0" by (auto simp add: Lim_cong_within \h \z\ z'\ \z'\0\ continuous_at h'_def) then obtain r2 where r2:"r2>0" "\x\ball z r2. h' x\0" using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto define r1 where "r1=min r2 r / 2" have "0 < r1" "cball z r1 \ ball z r" using \r2>0\ \r>0\ unfolding r1_def by auto moreover have "\x\cball z r1. h' x \ 0" using r2 unfolding r1_def by simp ultimately show ?thesis using that by auto qed then have "P h' 0 h' r1" using \h' holomorphic_on ball z r\ unfolding P_def by auto then have "P h 0 h' r1" unfolding P_def h'_def by auto then show ?thesis unfolding P_def by blast qed ultimately show ?thesis by auto qed have ?thesis when "\x. (f \ x) (at z)" apply (rule_tac imp_unique[unfolded P_def]) using P_exist[OF that(1) f_iso non_zero] unfolding P_def . moreover have ?thesis when "is_pole f z" proof (rule imp_unique[unfolded P_def]) obtain e where [simp]:"e>0" and e_holo:"f holomorphic_on ball z e - {z}" and e_nz: "\x\ball z e-{z}. f x\0" proof - have "\\<^sub>F z in at z. f z \ 0" using \is_pole f z\ filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def by auto then obtain e1 where e1:"e1>0" "\x\ball z e1-{z}. f x\0" using that eventually_at[of "\x. f x\0" z UNIV,simplified] by (auto simp add:dist_commute) obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto define e where "e=min e1 e2" show ?thesis apply (rule that[of e]) using e1 e2 unfolding e_def by auto qed define h where "h \ \x. inverse (f x)" have "\n g r. P h n g r" proof - have "h \z\ 0" using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce moreover have "\\<^sub>Fw in (at z). h w\0" using non_zero apply (elim frequently_rev_mp) unfolding h_def eventually_at by (auto intro:exI[where x=1]) moreover have "isolated_singularity_at h z" unfolding isolated_singularity_at_def h_def apply (rule exI[where x=e]) using e_holo e_nz \e>0\ by (metis open_ball analytic_on_open holomorphic_on_inverse open_delete) ultimately show ?thesis using P_exist[of h] by auto qed then obtain n g r where "0 < r" and g_holo:"g holomorphic_on cball z r" and "g z\0" and g_fac:"(\w\cball z r-{z}. h w = g w * (w - z) powr of_int n \ g w \ 0)" unfolding P_def by auto have "P f (-n) (inverse o g) r" proof - have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\cball z r - {z}" for w using g_fac[rule_format,of w] that unfolding h_def apply (auto simp add:powr_minus ) by (metis inverse_inverse_eq inverse_mult_distrib) then show ?thesis unfolding P_def comp_def using \r>0\ g_holo g_fac \g z\0\ by (auto intro:holomorphic_intros) qed then show "\x g r. 0 < r \ g holomorphic_on cball z r \ g z \ 0 \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int x \ g w \ 0)" unfolding P_def by blast qed ultimately show ?thesis using \not_essential f z\ unfolding not_essential_def by presburger qed lemma not_essential_transform: assumes "not_essential g z" assumes "\\<^sub>F w in (at z). g w = f w" shows "not_essential f z" using assms unfolding not_essential_def by (simp add: filterlim_cong is_pole_cong) lemma isolated_singularity_at_transform: assumes "isolated_singularity_at g z" assumes "\\<^sub>F w in (at z). g w = f w" shows "isolated_singularity_at f z" proof - obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}" using assms(1) unfolding isolated_singularity_at_def by auto obtain r2 where "r2>0" and r2:" \x. x \ z \ dist x z < r2 \ g x = f x" using assms(2) unfolding eventually_at by auto define r3 where "r3=min r1 r2" have "r3>0" unfolding r3_def using \r1>0\ \r2>0\ by auto moreover have "f analytic_on ball z r3 - {z}" proof - have "g holomorphic_on ball z r3 - {z}" using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto) then have "f holomorphic_on ball z r3 - {z}" using r2 unfolding r3_def by (auto simp add:dist_commute elim!:holomorphic_transform) then show ?thesis by (subst analytic_on_open,auto) qed ultimately show ?thesis unfolding isolated_singularity_at_def by auto qed lemma not_essential_powr[singularity_intros]: assumes "LIM w (at z). f w :> (at x)" shows "not_essential (\w. (f w) powr (of_int n)) z" proof - define fp where "fp=(\w. (f w) powr (of_int n))" have ?thesis when "n>0" proof - have "(\w. (f w) ^ (nat n)) \z\ x ^ nat n" using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros) then have "fp \z\ x ^ nat n" unfolding fp_def apply (elim Lim_transform_within[where d=1],simp) by (metis less_le powr_0 powr_of_int that zero_less_nat_eq zero_power) then show ?thesis unfolding not_essential_def fp_def by auto qed moreover have ?thesis when "n=0" proof - have "fp \z\ 1 " apply (subst tendsto_cong[where g="\_.1"]) using that filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def by auto then show ?thesis unfolding fp_def not_essential_def by auto qed moreover have ?thesis when "n<0" proof (cases "x=0") case True have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity" apply (subst filterlim_inverse_at_iff[symmetric],simp) apply (rule filterlim_atI) subgoal using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros) subgoal using filterlim_at_within_not_equal[OF assms,of 0] by (eventually_elim,insert that,auto) done then have "LIM w (at z). fp w :> at_infinity" proof (elim filterlim_mono_eventually) show "\\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x" using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def apply eventually_elim using powr_of_int that by auto qed auto then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto next case False let ?xx= "inverse (x ^ (nat (-n)))" have "(\w. inverse ((f w) ^ (nat (-n)))) \z\?xx" using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros) then have "fp \z\?xx" apply (elim Lim_transform_within[where d=1],simp) unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less not_le power_eq_0_iff powr_0 powr_of_int that) then show ?thesis unfolding fp_def not_essential_def by auto qed ultimately show ?thesis by linarith qed lemma isolated_singularity_at_powr[singularity_intros]: assumes "isolated_singularity_at f z" "\\<^sub>F w in (at z). f w\0" shows "isolated_singularity_at (\w. (f w) powr (of_int n)) z" proof - obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}" using assms(1) unfolding isolated_singularity_at_def by auto then have r1:"f holomorphic_on ball z r1 - {z}" using analytic_on_open[of "ball z r1-{z}" f] by blast obtain r2 where "r2>0" and r2:"\w. w \ z \ dist w z < r2 \ f w \ 0" using assms(2) unfolding eventually_at by auto define r3 where "r3=min r1 r2" have "(\w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}" apply (rule holomorphic_on_powr_of_int) subgoal unfolding r3_def using r1 by auto subgoal unfolding r3_def using r2 by (auto simp add:dist_commute) done moreover have "r3>0" unfolding r3_def using \0 < r1\ \0 < r2\ by linarith ultimately show ?thesis unfolding isolated_singularity_at_def apply (subst (asm) analytic_on_open[symmetric]) by auto qed lemma non_zero_neighbour: assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows "\\<^sub>F w in (at z). f w\0" proof - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto have "f w \ 0" when " w \ z" "dist w z < fr" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w \ 0" using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute) moreover have "(w - z) powr of_int fn \0" unfolding powr_eq_0_iff using \w\z\ by auto ultimately show ?thesis by auto qed then show ?thesis using \fr>0\ unfolding eventually_at by auto qed lemma non_zero_neighbour_pole: assumes "is_pole f z" shows "\\<^sub>F w in (at z). f w\0" using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0] unfolding is_pole_def by auto lemma non_zero_neighbour_alt: assumes holo: "f holomorphic_on S" and "open S" "connected S" "z \ S" "\ \ S" "f \ \ 0" shows "\\<^sub>F w in (at z). f w\0 \ w\S" proof (cases "f z = 0") case True from isolated_zeros[OF holo \open S\ \connected S\ \z \ S\ True \\ \ S\ \f \ \ 0\] obtain r where "0 < r" "ball z r \ S" "\w \ ball z r - {z}.f w \ 0" by metis then show ?thesis unfolding eventually_at apply (rule_tac x=r in exI) by (auto simp add:dist_commute) next case False obtain r1 where r1:"r1>0" "\y. dist z y < r1 \ f y \ 0" using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at holo holomorphic_on_imp_continuous_on by blast obtain r2 where r2:"r2>0" "ball z r2 \ S" using assms(2) assms(4) openE by blast show ?thesis unfolding eventually_at apply (rule_tac x="min r1 r2" in exI) using r1 r2 by (auto simp add:dist_commute) qed lemma not_essential_times[singularity_intros]: assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "not_essential (\w. f w * g w) z" proof - define fg where "fg = (\w. f w * g w)" have ?thesis when "\ ((\\<^sub>Fw in (at z). f w\0) \ (\\<^sub>Fw in (at z). g w\0))" proof - have "\\<^sub>Fw in (at z). fg w=0" using that[unfolded frequently_def, simplified] unfolding fg_def by (auto elim: eventually_rev_mp) from tendsto_cong[OF this] have "fg \z\0" by auto then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" and g_nconst:"\\<^sub>Fw in (at z). g w\0" proof - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto obtain gn gp gr where [simp]:"gp z \ 0" and "gr > 0" and gr: "gp holomorphic_on cball z gr" "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto define r1 where "r1=(min fr gr)" have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" when "w\ball z r1 - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that unfolding r1_def by auto moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" using gr(2)[rule_format, of w] that unfolding r1_def by auto ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" unfolding fg_def by (auto simp add:powr_add) qed have [intro]: "fp \z\fp z" "gp \z\gp z" using fr(1) \fr>0\ gr(1) \gr>0\ by (meson open_ball ball_subset_cball centre_in_ball continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on holomorphic_on_subset)+ have ?thesis when "fn+gn>0" proof - have "(\w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \z\0" using that by (auto intro!:tendsto_eq_intros) then have "fg \z\ 0" apply (elim Lim_transform_within[OF _ \r1>0\]) by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int that) then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when "fn+gn=0" proof - have "(\w. fp w * gp w) \z\fp z*gp z" using that by (auto intro!:tendsto_eq_intros) then have "fg \z\ fp z*gp z" apply (elim Lim_transform_within[OF _ \r1>0\]) apply (subst fg_times) by (auto simp add:dist_commute that) then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when "fn+gn<0" proof - have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity" apply (rule filterlim_divide_at_infinity) apply (insert that, auto intro!:tendsto_eq_intros filterlim_atI) using eventually_at_topological by blast then have "is_pole fg z" unfolding is_pole_def apply (elim filterlim_transform_within[OF _ _ \r1>0\],simp) apply (subst fg_times,simp add:dist_commute) apply (subst powr_of_int) using that by (auto simp add:field_split_simps) then show ?thesis unfolding not_essential_def fg_def by auto qed ultimately show ?thesis unfolding not_essential_def fg_def by fastforce qed ultimately show ?thesis by auto qed lemma not_essential_inverse[singularity_intros]: assumes f_ness:"not_essential f z" assumes f_iso:"isolated_singularity_at f z" shows "not_essential (\w. inverse (f w)) z" proof - define vf where "vf = (\w. inverse (f w))" have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" proof - have "\\<^sub>Fw in (at z). f w=0" using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) then have "\\<^sub>Fw in (at z). vf w=0" unfolding vf_def by auto from tendsto_cong[OF this] have "vf \z\0" unfolding vf_def by auto then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "is_pole f z" proof - have "vf \z\0" using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "\x. f\z\x " and f_nconst:"\\<^sub>Fw in (at z). f w\0" proof - from that obtain fz where fz:"f\z\fz" by auto have ?thesis when "fz=0" proof - have "(\w. inverse (vf w)) \z\0" using fz that unfolding vf_def by auto moreover have "\\<^sub>F w in at z. inverse (vf w) \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] unfolding vf_def by auto ultimately have "is_pole vf z" using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "fz\0" proof - have "vf \z\inverse fz" using fz that unfolding vf_def by (auto intro:tendsto_eq_intros) then show ?thesis unfolding not_essential_def vf_def by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis using f_ness unfolding not_essential_def by auto qed lemma isolated_singularity_at_inverse[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" shows "isolated_singularity_at (\w. inverse (f w)) z" proof - define vf where "vf = (\w. inverse (f w))" have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" proof - have "\\<^sub>Fw in (at z). f w=0" using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) then have "\\<^sub>Fw in (at z). vf w=0" unfolding vf_def by auto then obtain d1 where "d1>0" and d1:"\x. x \ z \ dist x z < d1 \ vf x = 0" unfolding eventually_at by auto then have "vf holomorphic_on ball z d1-{z}" apply (rule_tac holomorphic_transform[of "\_. 0"]) by (auto simp add:dist_commute) then have "vf analytic_on ball z d1 - {z}" by (simp add: analytic_on_open open_delete) then show ?thesis using \d1>0\ unfolding isolated_singularity_at_def vf_def by auto qed moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" proof - have "\\<^sub>F w in at z. f w \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] . then obtain d1 where d1:"d1>0" "\x. x \ z \ dist x z < d1 \ f x \ 0" unfolding eventually_at by auto obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}" using f_iso unfolding isolated_singularity_at_def by auto define d3 where "d3=min d1 d2" have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto moreover have "vf analytic_on ball z d3 - {z}" unfolding vf_def apply (rule analytic_on_inverse) subgoal using d2 unfolding d3_def by (elim analytic_on_subset) auto subgoal for w using d1 unfolding d3_def by (auto simp add:dist_commute) done ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto qed ultimately show ?thesis by auto qed lemma not_essential_divide[singularity_intros]: assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "not_essential (\w. f w / g w) z" proof - have "not_essential (\w. f w * inverse (g w)) z" apply (rule not_essential_times[where g="\w. inverse (g w)"]) using assms by (auto intro: isolated_singularity_at_inverse not_essential_inverse) then show ?thesis by (simp add:field_simps) qed lemma assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows isolated_singularity_at_times[singularity_intros]: "isolated_singularity_at (\w. f w * g w) z" and isolated_singularity_at_add[singularity_intros]: "isolated_singularity_at (\w. f w + g w) z" proof - obtain d1 d2 where "d1>0" "d2>0" and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}" using f_iso g_iso unfolding isolated_singularity_at_def by auto define d3 where "d3=min d1 d2" have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto have "(\w. f w * g w) analytic_on ball z d3 - {z}" apply (rule analytic_on_mult) using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) then show "isolated_singularity_at (\w. f w * g w) z" using \d3>0\ unfolding isolated_singularity_at_def by auto have "(\w. f w + g w) analytic_on ball z d3 - {z}" apply (rule analytic_on_add) using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) then show "isolated_singularity_at (\w. f w + g w) z" using \d3>0\ unfolding isolated_singularity_at_def by auto qed lemma isolated_singularity_at_uminus[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" shows "isolated_singularity_at (\w. - f w) z" using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast lemma isolated_singularity_at_id[singularity_intros]: "isolated_singularity_at (\w. w) z" unfolding isolated_singularity_at_def by (simp add: gt_ex) lemma isolated_singularity_at_minus[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "isolated_singularity_at (\w. f w - g w) z" using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\w. - g w"] ,OF g_iso] by simp lemma isolated_singularity_at_divide[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and g_ness:"not_essential g z" shows "isolated_singularity_at (\w. f w / g w) z" using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso, of "\w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps) lemma isolated_singularity_at_const[singularity_intros]: "isolated_singularity_at (\w. c) z" unfolding isolated_singularity_at_def by (simp add: gt_ex) lemma isolated_singularity_at_holomorphic: assumes "f holomorphic_on s-{z}" "open s" "z\s" shows "isolated_singularity_at f z" using assms unfolding isolated_singularity_at_def by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) subsubsection \The order of non-essential singularities (i.e. removable singularities or poles)\ definition\<^marker>\tag important\ zorder :: "(complex \ complex) \ complex \ int" where "zorder f z = (THE n. (\h r. r>0 \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. f w = h w * (w-z) powr (of_int n) \ h w \0)))" definition\<^marker>\tag important\ zor_poly ::"[complex \ complex, complex] \ complex \ complex" where "zor_poly f z = (SOME h. \r. r > 0 \ h holomorphic_on cball z r \ h z \ 0 \ (\w\cball z r - {z}. f w = h w * (w - z) powr (zorder f z) \ h w \0))" lemma zorder_exist: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows "g z\0 \ (\r. r>0 \ g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w-z) powr n \ g w \0))" proof - define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have "\!n. \g r. P n g r" using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto then have "\g r. P n g r" unfolding n_def P_def zorder_def by (drule_tac theI',argo) then have "\r. P n g r" unfolding P_def zor_poly_def g_def n_def by (drule_tac someI_ex,argo) then obtain r1 where "P n g r1" by auto then show ?thesis unfolding P_def by auto qed lemma fixes f::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows zorder_inverse: "zorder (\w. inverse (f w)) z = - zorder f z" and zor_poly_inverse: "\\<^sub>Fw in (at z). zor_poly (\w. inverse (f w)) z w = inverse (zor_poly f z w)" proof - define vf where "vf = (\w. inverse (f w))" define fn vfn where "fn = zorder f z" and "vfn = zorder vf z" define fp vfp where "fp = zor_poly f z" and "vfp = zor_poly vf z" obtain fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def] by auto have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" and fr_nz: "inverse (fp w)\0" when "w\ball z fr - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that by auto then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\0" unfolding vf_def by (auto simp add:powr_minus) qed obtain vfr where [simp]:"vfp z \ 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr" "(\w\cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0)" proof - have "isolated_singularity_at vf z" using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def . moreover have "not_essential vf z" using not_essential_inverse[OF f_ness f_iso] unfolding vf_def . moreover have "\\<^sub>F w in at z. vf w \ 0" using f_nconst unfolding vf_def by (auto elim:frequently_elim1) ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto qed define r1 where "r1 = min fr vfr" have "r1>0" using \fr>0\ \vfr>0\ unfolding r1_def by simp show "vfn = - fn" apply (rule holomorphic_factor_unique[of r1 vfp z "\w. inverse (fp w)" vf]) subgoal using \r1>0\ by simp subgoal by simp subgoal by simp subgoal proof (rule ballI) fix w assume "w \ ball z r1 - {z}" then have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" unfolding r1_def by auto from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] show "vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0 \ vf w = inverse (fp w) * (w - z) powr of_int (- fn) \ inverse (fp w) \ 0" by auto qed subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros) subgoal using fr unfolding r1_def by (auto intro!:holomorphic_intros) done have "vfp w = inverse (fp w)" when "w\ball z r1-{z}" for w proof - have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" "w\z" using that unfolding r1_def by auto from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \vfn = - fn\ \w\z\ show ?thesis by auto qed then show "\\<^sub>Fw in (at z). vfp w = inverse (fp w)" unfolding eventually_at using \r1>0\ apply (rule_tac x=r1 in exI) by (auto simp add:dist_commute) qed lemma fixes f g::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and f_ness:"not_essential f z" and g_ness:"not_essential g z" and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" shows zorder_times:"zorder (\w. f w * g w) z = zorder f z + zorder g z" and zor_poly_times:"\\<^sub>Fw in (at z). zor_poly (\w. f w * g w) z w = zor_poly f z w *zor_poly g z w" proof - define fg where "fg = (\w. f w * g w)" define fn gn fgn where "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z" define fp gp fgp where "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z" have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" using fg_nconst by (auto elim!:frequently_elim1) obtain fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto obtain gr where [simp]:"gp z \ 0" and "gr > 0" and gr: "gp holomorphic_on cball z gr" "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto define r1 where "r1=min fr gr" have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" when "w\ball z r1 - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that unfolding r1_def by auto moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" using gr(2)[rule_format, of w] that unfolding r1_def by auto ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" unfolding fg_def by (auto simp add:powr_add) qed obtain fgr where [simp]:"fgp z \ 0" and "fgr > 0" and fgr: "fgp holomorphic_on cball z fgr" "\w\cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0" proof - have "fgp z \ 0 \ (\r>0. fgp holomorphic_on cball z r \ (\w\cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0))" apply (rule zorder_exist[of fg z, folded fgn_def fgp_def]) subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] . subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] . subgoal unfolding fg_def using fg_nconst . done then show ?thesis using that by blast qed define r2 where "r2 = min fgr r1" have "r2>0" using \r1>0\ \fgr>0\ unfolding r2_def by simp show "fgn = fn + gn " apply (rule holomorphic_factor_unique[of r2 fgp z "\w. fp w * gp w" fg]) subgoal using \r2>0\ by simp subgoal by simp subgoal by simp subgoal proof (rule ballI) fix w assume "w \ ball z r2 - {z}" then have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" unfolding r2_def by auto from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] show "fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0 \ fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \ fp w * gp w \ 0" by auto qed subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) done have "fgp w = fp w *gp w" when "w\ball z r2-{z}" for w proof - have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" "w\z" using that unfolding r2_def by auto from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \fgn = fn + gn\ \w\z\ show ?thesis by auto qed then show "\\<^sub>Fw in (at z). fgp w = fp w * gp w" using \r2>0\ unfolding eventually_at by (auto simp add:dist_commute) qed lemma fixes f g::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and f_ness:"not_essential f z" and g_ness:"not_essential g z" and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" shows zorder_divide:"zorder (\w. f w / g w) z = zorder f z - zorder g z" and zor_poly_divide:"\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" proof - have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" using fg_nconst by (auto elim!:frequently_elim1) define vg where "vg=(\w. inverse (g w))" have "zorder (\w. f w * vg w) z = zorder f z + zorder vg z" apply (rule zorder_times[OF f_iso _ f_ness,of vg]) subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) done then show "zorder (\w. f w / g w) z = zorder f z - zorder g z" using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def by (auto simp add:field_simps) have "\\<^sub>F w in at z. zor_poly (\w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w" apply (rule zor_poly_times[OF f_iso _ f_ness,of vg]) subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) done then show "\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def apply eventually_elim by (auto simp add:field_simps) qed lemma zorder_exist_zero: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes holo: "f holomorphic_on s" and "open s" "connected s" "z\s" and non_const: "\w\s. f w \ 0" shows "(if f z=0 then n > 0 else n=0) \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r \ (\w\cball z r. f w = g w * (w-z) ^ nat n \ g w \0))" proof - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" proof - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" proof (rule zorder_exist[of f z,folded g_def n_def]) show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using holo assms(4,6) by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE) show "not_essential f z" unfolding not_essential_def using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce have "\\<^sub>F w in at z. f w \ 0 \ w\s" proof - obtain w where "w\s" "f w\0" using non_const by auto then show ?thesis by (rule non_zero_neighbour_alt[OF holo \open s\ \connected s\ \z\s\]) qed then show "\\<^sub>F w in at z. f w \ 0" apply (elim eventually_frequentlyE) by auto qed then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" by auto obtain r2 where r2: "r2>0" "cball z r2 \ s" using assms(4,6) open_contains_cball_eq by blast define r3 where "r3=min r1 r2" have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto moreover have "g holomorphic_on cball z r3" using r1(1) unfolding r3_def by auto moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" using r1(2) unfolding r3_def by auto ultimately show ?thesis using that[of r3] \g z\0\ by auto qed have if_0:"if f z=0 then n > 0 else n=0" proof - have "f\ z \ f z" by (metis assms(4,6,7) at_within_open continuous_on holo holomorphic_on_imp_continuous_on) then have "(\w. g w * (w - z) powr of_int n) \z\ f z" apply (elim Lim_transform_within_open[where s="ball z r"]) using r by auto moreover have "g \z\g z" by (metis (mono_tags, lifting) open_ball at_within_open_subset ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1,3) subsetCE) ultimately have "(\w. (g w * (w - z) powr of_int n) / g w) \z\ f z/g z" apply (rule_tac tendsto_divide) using \g z\0\ by auto then have powr_tendsto:"(\w. (w - z) powr of_int n) \z\ f z/g z" apply (elim Lim_transform_within_open[where s="ball z r"]) using r by auto have ?thesis when "n\0" "f z=0" proof - have "(\w. (w - z) ^ nat n) \z\ f z/g z" using powr_tendsto apply (elim Lim_transform_within[where d=r]) by (auto simp add: powr_of_int \n\0\ \r>0\) then have *:"(\w. (w - z) ^ nat n) \z\ 0" using \f z=0\ by simp moreover have False when "n=0" proof - have "(\w. (w - z) ^ nat n) \z\ 1" using \n=0\ by auto then show False using * using LIM_unique zero_neq_one by blast qed ultimately show ?thesis using that by fastforce qed moreover have ?thesis when "n\0" "f z\0" proof - have False when "n>0" proof - have "(\w. (w - z) ^ nat n) \z\ f z/g z" using powr_tendsto apply (elim Lim_transform_within[where d=r]) by (auto simp add: powr_of_int \n\0\ \r>0\) moreover have "(\w. (w - z) ^ nat n) \z\ 0" using \n>0\ by (auto intro!:tendsto_eq_intros) ultimately show False using \f z\0\ \g z\0\ using LIM_unique divide_eq_0_iff by blast qed then show ?thesis using that by force qed moreover have False when "n<0" proof - have "(\w. inverse ((w - z) ^ nat (- n))) \z\ f z/g z" "(\w.((w - z) ^ nat (- n))) \z\ 0" subgoal using powr_tendsto powr_of_int that by (elim Lim_transform_within_open[where s=UNIV],auto) subgoal using that by (auto intro!:tendsto_eq_intros) done from tendsto_mult[OF this,simplified] have "(\x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \z\ 0" . then have "(\x. 1::complex) \z\ 0" by (elim Lim_transform_within_open[where s=UNIV],auto) then show False using LIM_const_eq by fastforce qed ultimately show ?thesis by fastforce qed moreover have "f w = g w * (w-z) ^ nat n \ g w \0" when "w\cball z r" for w proof (cases "w=z") case True then have "f \z\f w" using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce then have "(\w. g w * (w-z) ^ nat n) \z\f w" proof (elim Lim_transform_within[OF _ \r>0\]) fix x assume "0 < dist x z" "dist x z < r" then have "x \ cball z r - {z}" "x\z" unfolding cball_def by (auto simp add: dist_commute) then have "f x = g x * (x - z) powr of_int n" using r(4)[rule_format,of x] by simp also have "... = g x * (x - z) ^ nat n" apply (subst powr_of_int) using if_0 \x\z\ by (auto split:if_splits) finally show "f x = g x * (x - z) ^ nat n" . qed moreover have "(\w. g w * (w-z) ^ nat n) \z\ g w * (w-z) ^ nat n" using True apply (auto intro!:tendsto_eq_intros) by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1) r(3) that) ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast then show ?thesis using \g z\0\ True by auto next case False then have "f w = g w * (w - z) powr of_int n \ g w \ 0" using r(4) that by auto then show ?thesis using False if_0 powr_of_int by (auto split:if_splits) qed ultimately show ?thesis using r by auto qed lemma zorder_exist_pole: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes holo: "f holomorphic_on s-{z}" and "open s" "z\s" and "is_pole f z" shows "n < 0 \ g z\0 \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0))" proof - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" proof - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" proof (rule zorder_exist[of f z,folded g_def n_def]) show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using holo assms(4,5) by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) show "not_essential f z" unfolding not_essential_def using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce from non_zero_neighbour_pole[OF \is_pole f z\] show "\\<^sub>F w in at z. f w \ 0" apply (elim eventually_frequentlyE) by auto qed then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" by auto obtain r2 where r2: "r2>0" "cball z r2 \ s" using assms(4,5) open_contains_cball_eq by metis define r3 where "r3=min r1 r2" have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto moreover have "g holomorphic_on cball z r3" using r1(1) unfolding r3_def by auto moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" using r1(2) unfolding r3_def by auto ultimately show ?thesis using that[of r3] \g z\0\ by auto qed have "n<0" proof (rule ccontr) assume " \ n < 0" define c where "c=(if n=0 then g z else 0)" have [simp]:"g \z\ g z" by (metis open_ball at_within_open ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) ) have "\\<^sub>F x in at z. f x = g x * (x - z) ^ nat n" unfolding eventually_at_topological apply (rule_tac exI[where x="ball z r"]) using r powr_of_int \\ n < 0\ by auto moreover have "(\x. g x * (x - z) ^ nat n) \z\c" proof (cases "n=0") case True then show ?thesis unfolding c_def by simp next case False then have "(\x. (x - z) ^ nat n) \z\ 0" using \\ n < 0\ by (auto intro!:tendsto_eq_intros) from tendsto_mult[OF _ this,of g "g z",simplified] show ?thesis unfolding c_def using False by simp qed ultimately have "f \z\c" using tendsto_cong by fast then show False using \is_pole f z\ at_neq_bot not_tendsto_and_filterlim_at_infinity unfolding is_pole_def by blast qed moreover have "\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0" using r(4) \n<0\ powr_of_int by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le) ultimately show ?thesis using r(1-3) \g z\0\ by auto qed lemma zorder_eqI: assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" assumes fg_eq:"\w. \w \ s;w\z\ \ f w = g w * (w - z) powr n" shows "zorder f z = n" proof - have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact moreover have "open (-{0::complex})" by auto ultimately have "open ((g -` (-{0})) \ s)" unfolding continuous_on_open_vimage[OF \open s\] by blast moreover from assms have "z \ (g -` (-{0})) \ s" by auto ultimately obtain r where r: "r > 0" "cball z r \ s \ (g -` (-{0}))" unfolding open_contains_cball by blast let ?gg= "(\w. g w * (w - z) powr n)" define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have "P n g r" unfolding P_def using r assms(3,4,5) by auto then have "\g r. P n g r" by auto moreover have unique: "\!n. \g r. P n g r" unfolding P_def proof (rule holomorphic_factor_puncture) have "ball z r-{z} \ s" using r using ball_subset_cball by blast then have "?gg holomorphic_on ball z r-{z}" using \g holomorphic_on s\ r by (auto intro!: holomorphic_intros) then have "f holomorphic_on ball z r - {z}" apply (elim holomorphic_transform) using fg_eq \ball z r-{z} \ s\ by auto then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using analytic_on_open open_delete r(1) by blast next have "not_essential ?gg z" proof (intro singularity_intros) show "not_essential g z" by (meson \continuous_on s g\ assms(1) assms(2) continuous_on_eq_continuous_at isCont_def not_essential_def) show " \\<^sub>F w in at z. w - z \ 0" by (simp add: eventually_at_filter) then show "LIM w at z. w - z :> at 0" unfolding filterlim_at by (auto intro:tendsto_eq_intros) show "isolated_singularity_at g z" by (meson Diff_subset open_ball analytic_on_holomorphic assms(1,2,3) holomorphic_on_subset isolated_singularity_at_def openE) qed then show "not_essential f z" apply (elim not_essential_transform) unfolding eventually_at using assms(1,2) assms(5)[symmetric] by (metis dist_commute mem_ball openE subsetCE) show "\\<^sub>F w in at z. f w \ 0" unfolding frequently_at proof (rule,rule) fix d::real assume "0 < d" define z' where "z'=z+min d r / 2" have "z' \ z" " dist z' z < d " unfolding z'_def using \d>0\ \r>0\ by (auto simp add:dist_norm) moreover have "f z' \ 0" proof (subst fg_eq[OF _ \z'\z\]) have "z' \ cball z r" unfolding z'_def using \r>0\ \d>0\ by (auto simp add:dist_norm) then show " z' \ s" using r(2) by blast show "g z' * (z' - z) powr of_int n \ 0" using P_def \P n g r\ \z' \ cball z r\ calculation(1) by auto qed ultimately show "\x\UNIV. x \ z \ dist x z < d \ f x \ 0" by auto qed qed ultimately have "(THE n. \g r. P n g r) = n" by (rule_tac the1_equality) then show ?thesis unfolding zorder_def P_def by blast qed lemma residue_pole_order: fixes f::"complex \ complex" and z::complex defines "n \ nat (- zorder f z)" and "h \ zor_poly f z" assumes f_iso:"isolated_singularity_at f z" and pole:"is_pole f z" shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ ball z e" and h_holo: "h holomorphic_on cball z r" and h_divide:"(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" proof - obtain r where r:"zorder f z < 0" "h z \ 0" "r>0" "cball z r \ ball z e" "h holomorphic_on cball z r" "(\w\cball z r - {z}. f w = h w / (w - z) ^ n \ h w \ 0)" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\,folded n_def h_def] by auto have "n>0" using \zorder f z < 0\ unfolding n_def by simp moreover have "(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" using \h z\0\ r(6) by blast ultimately show ?thesis using r(3,4,5) that by blast qed have r_nonzero:"\w. w \ ball z r - {z} \ f w \ 0" using h_divide by simp define c where "c \ 2 * pi * \" define der_f where "der_f \ ((deriv ^^ (n - 1)) h z / fact (n-1))" define h' where "h' \ \u. h u / (u - z) ^ n" have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" unfolding h'_def proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", folded c_def Suc_pred'[OF \n>0\]]) show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp show "h holomorphic_on ball z r" using h_holo by auto show " z \ ball z r" using \r>0\ by auto qed then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto then have "(f has_contour_integral c * der_f) (circlepath z r)" proof (elim has_contour_integral_eq) fix x assume "x \ path_image (circlepath z r)" hence "x\cball z r - {z}" using \r>0\ by auto then show "h' x = f x" using h_divide unfolding h'_def by auto qed moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" using base_residue[of \ball z e\ z,simplified,OF \r>0\ f_holo r_cball,folded c_def] unfolding c_def by simp ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast hence "der_f = residue f z" unfolding c_def by auto thus ?thesis unfolding der_f_def by auto qed lemma simple_zeroI: assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" assumes "\w. w \ s \ f w = g w * (w - z)" shows "zorder f z = 1" using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto) lemma higher_deriv_power: shows "(deriv ^^ j) (\w. (w - z) ^ n) w = pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)" proof (induction j arbitrary: w) case 0 thus ?case by auto next case (Suc j w) have "(deriv ^^ Suc j) (\w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\w. (w - z) ^ n)) w" by simp also have "(deriv ^^ j) (\w. (w - z) ^ n) = (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))" using Suc by (intro Suc.IH ext) also { have "(\ has_field_derivative of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)" using Suc.prems by (auto intro!: derivative_eq_intros) also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = pochhammer (of_nat (Suc n - Suc j)) (Suc j)" by (cases "Suc j \ n", subst pochhammer_rec) (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left) finally have "deriv (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w = \ * (w - z) ^ (n - Suc j)" by (rule DERIV_imp_deriv) } finally show ?case . qed lemma zorder_zero_eqI: assumes f_holo:"f holomorphic_on s" and "open s" "z \ s" assumes zero: "\i. i < nat n \ (deriv ^^ i) f z = 0" assumes nz: "(deriv ^^ nat n) f z \ 0" and "n\0" shows "zorder f z = n" proof - obtain r where [simp]:"r>0" and "ball z r \ s" using \open s\ \z\s\ openE by blast have nz':"\w\ball z r. f w \ 0" proof (rule ccontr) assume "\ (\w\ball z r. f w \ 0)" then have "eventually (\u. f u = 0) (nhds z)" using \r>0\ unfolding eventually_nhds apply (rule_tac x="ball z r" in exI) by auto then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\_. 0) z" by (intro higher_deriv_cong_ev) auto also have "(deriv ^^ nat n) (\_. 0) z = 0" by (induction n) simp_all finally show False using nz by contradiction qed define zn g where "zn = zorder f z" and "g = zor_poly f z" obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and [simp]:"e>0" and "cball z e \ ball z r" and g_holo:"g holomorphic_on cball z e" and e_fac:"(\w\cball z e. f w = g w * (w - z) ^ nat zn \ g w \ 0)" proof - have "f holomorphic_on ball z r" using f_holo \ball z r \ s\ by auto from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def] show ?thesis by blast qed from this(1,2,5) have "zn\0" "g z\0" subgoal by (auto split:if_splits) subgoal using \0 < e\ ball_subset_cball centre_in_ball e_fac by blast done define A where "A = (\i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)" have deriv_A:"(deriv ^^ i) f z = (if zn \ int i then A i else 0)" for i proof - have "eventually (\w. w \ ball z e) (nhds z)" using \cball z e \ ball z r\ \e>0\ by (intro eventually_nhds_in_open) auto hence "eventually (\w. f w = (w - z) ^ (nat zn) * g w) (nhds z)" apply eventually_elim by (use e_fac in auto) hence "(deriv ^^ i) f z = (deriv ^^ i) (\w. (w - z) ^ nat zn * g w) z" by (intro higher_deriv_cong_ev) auto also have "\ = (\j=0..i. of_nat (i choose j) * (deriv ^^ j) (\w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)" using g_holo \e>0\ by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros) also have "\ = (\j=0..i. if j = nat zn then of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" proof (intro sum.cong refl, goal_cases) case (1 j) have "(deriv ^^ j) (\w. (w - z) ^ nat zn) z = pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)" by (subst higher_deriv_power) auto also have "\ = (if j = nat zn then fact j else 0)" by (auto simp: not_less pochhammer_0_left pochhammer_fact) also have "of_nat (i choose j) * \ * (deriv ^^ (i - j)) g z = (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" by simp finally show ?case . qed also have "\ = (if i \ zn then A i else 0)" by (auto simp: A_def) finally show "(deriv ^^ i) f z = \" . qed have False when "nn\0\ by auto with nz show False by auto qed moreover have "n\zn" proof - have "g z \ 0" using e_fac[rule_format,of z] \e>0\ by simp then have "(deriv ^^ nat zn) f z \ 0" using deriv_A[of "nat zn"] by(auto simp add:A_def) then have "nat zn \ nat n" using zero[of "nat zn"] by linarith moreover have "zn\0" using e_if by (auto split:if_splits) ultimately show ?thesis using nat_le_eq_zle by blast qed ultimately show ?thesis unfolding zn_def by fastforce qed lemma assumes "eventually (\z. f z = g z) (at z)" "z = z'" shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'" proof - define P where "P = (\ff n h r. 0 < r \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \ h w\0))" have "(\r. P f n h r) = (\r. P g n h r)" for n h proof - have *: "\r. P g n h r" if "\r. P f n h r" and "eventually (\x. f x = g x) (at z)" for f g proof - from that(1) obtain r1 where r1_P:"P f n h r1" by auto from that(2) obtain r2 where "r2>0" and r2_dist:"\x. x \ z \ dist x z \ r2 \ f x = g x" unfolding eventually_at_le by auto define r where "r=min r1 r2" have "r>0" "h z\0" using r1_P \r2>0\ unfolding r_def P_def by auto moreover have "h holomorphic_on cball z r" using r1_P unfolding P_def r_def by auto moreover have "g w = h w * (w - z) powr of_int n \ h w \ 0" when "w\cball z r - {z}" for w proof - have "f w = h w * (w - z) powr of_int n \ h w \ 0" using r1_P that unfolding P_def r_def by auto moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def by (simp add: dist_commute) ultimately show ?thesis by simp qed ultimately show ?thesis unfolding P_def by auto qed from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) show ?thesis by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']]) qed then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'" using \z=z'\ unfolding P_def zorder_def zor_poly_def by auto qed lemma zorder_nonzero_div_power: assumes "open s" "z \ s" "f holomorphic_on s" "f z \ 0" "n > 0" shows "zorder (\w. f w / (w - z) ^ n) z = - n" apply (rule zorder_eqI[OF \open s\ \z\s\ \f holomorphic_on s\ \f z\0\]) apply (subst powr_of_int) using \n>0\ by (auto simp add:field_simps) lemma zor_poly_eq: assumes "isolated_singularity_at f z" "not_essential f z" "\\<^sub>F w in at z. f w \ 0" shows "eventually (\w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)" proof - obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))" using zorder_exist[OF assms] by blast then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z" by (auto simp: field_simps powr_minus) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_zero_eq: assumes "f holomorphic_on s" "open s" "connected s" "z \ s" "\w\s. f w \ 0" shows "eventually (\w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)" proof - obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))" using zorder_exist_zero[OF assms] by auto then have *: "\w\ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)" by (auto simp: field_simps powr_minus) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_pole_eq: assumes f_iso:"isolated_singularity_at f z" "is_pole f z" shows "eventually (\w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)" proof - obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\] by auto then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)" by (auto simp: field_simps) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes "isolated_singularity_at f z0" "not_essential f z0" "\\<^sub>F w in at z0. f w \ 0" assumes lim: "((\x. f (g x) * (g x - z0) powr - n) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - from zorder_exist[OF assms(2-4)] obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" "\w. w \ cball z0 r - {z0} \ f w = zor_poly f z0 w * (w - z0) powr n" unfolding n_def by blast from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto hence "eventually (\w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)" by eventually_elim (insert r, auto simp: field_simps powr_minus) moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w * (w - z0) powr - n) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) * (g x - z0) powr - n) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma zor_poly_zero_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes "f holomorphic_on A" "open A" "connected A" "z0 \ A" "\z\A. f z \ 0" assumes lim: "((\x. f (g x) / (g x - z0) ^ nat n) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - from zorder_exist_zero[OF assms(2-6)] obtain r where r: "r > 0" "cball z0 r \ A" "zor_poly f z0 holomorphic_on cball z0 r" "\w. w \ cball z0 r \ f w = zor_poly f z0 w * (w - z0) ^ nat n" unfolding n_def by blast from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto hence "eventually (\w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)" by eventually_elim (insert r, auto simp: field_simps) moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w / (w - z0) ^ nat n) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) / (g x - z0) ^ nat n) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma zor_poly_pole_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0" assumes lim: "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" proof - have "\\<^sub>F w in at z0. f w \ 0" using non_zero_neighbour_pole[OF \is_pole f z0\] by (auto elim:eventually_frequentlyE) moreover have "not_essential f z0" unfolding not_essential_def using \is_pole f z0\ by simp ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto qed from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto have "eventually (\w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)" using zor_poly_pole_eq[OF f_iso \is_pole f z0\] unfolding n_def . moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w * (w - z0) ^ nat (-n)) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma residue_simple_pole: assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" shows "residue f z0 = zor_poly f z0 z0" using assms by (subst residue_pole_order) simp_all lemma residue_simple_pole_limit: assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" assumes "((\x. f (g x) * (g x - z0)) \ c) F" assumes "filterlim g (at z0) F" "F \ bot" shows "residue f z0 = c" proof - have "residue f z0 = zor_poly f z0 z0" by (rule residue_simple_pole assms)+ also have "\ = c" apply (rule zor_poly_pole_eqI) using assms by auto finally show ?thesis . qed lemma lhopital_complex_simple: assumes "(f has_field_derivative f') (at z)" assumes "(g has_field_derivative g') (at z)" assumes "f z = 0" "g z = 0" "g' \ 0" "f' / g' = c" shows "((\w. f w / g w) \ c) (at z)" proof - have "eventually (\w. w \ z) (at z)" by (auto simp: eventually_at_filter) hence "eventually (\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)" by eventually_elim (simp add: assms field_split_simps) moreover have "((\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \ f' / g') (at z)" by (intro tendsto_divide has_field_derivativeD assms) ultimately have "((\w. f w / g w) \ f' / g') (at z)" by (blast intro: Lim_transform_eventually) with assms show ?thesis by simp qed lemma assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" and "open s" "connected s" "z \ s" assumes g_deriv:"(g has_field_derivative g') (at z)" assumes "f z \ 0" "g z = 0" "g' \ 0" shows porder_simple_pole_deriv: "zorder (\w. f w / g w) z = - 1" and residue_simple_pole_deriv: "residue (\w. f w / g w) z = f z / g'" proof - have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z" using isolated_singularity_at_holomorphic[OF _ \open s\ \z\s\] f_holo g_holo by (meson Diff_subset holomorphic_on_subset)+ have [simp]:"not_essential f z" "not_essential g z" unfolding not_essential_def using f_holo g_holo assms(3,5) by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+ have g_nconst:"\\<^sub>F w in at z. g w \0 " proof (rule ccontr) assume "\ (\\<^sub>F w in at z. g w \ 0)" then have "\\<^sub>F w in nhds z. g w = 0" unfolding eventually_at eventually_nhds frequently_at using \g z = 0\ by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball) then have "deriv g z = deriv (\_. 0) z" by (intro deriv_cong_ev) auto then have "deriv g z = 0" by auto then have "g' = 0" using g_deriv DERIV_imp_deriv by blast then show False using \g'\0\ by auto qed have "zorder (\w. f w / g w) z = zorder f z - zorder g z" proof - have "\\<^sub>F w in at z. f w \0 \ w\s" apply (rule non_zero_neighbour_alt) using assms by auto with g_nconst have "\\<^sub>F w in at z. f w * g w \ 0" by (elim frequently_rev_mp eventually_rev_mp,auto) then show ?thesis using zorder_divide[of f z g] by auto qed moreover have "zorder f z=0" apply (rule zorder_zero_eqI[OF f_holo \open s\ \z\s\]) using \f z\0\ by auto moreover have "zorder g z=1" apply (rule zorder_zero_eqI[OF g_holo \open s\ \z\s\]) subgoal using assms(8) by auto subgoal using DERIV_imp_deriv assms(9) g_deriv by auto subgoal by simp done ultimately show "zorder (\w. f w / g w) z = - 1" by auto show "residue (\w. f w / g w) z = f z / g'" proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified]) show "zorder (\w. f w / g w) z = - 1" by fact show "isolated_singularity_at (\w. f w / g w) z" by (auto intro: singularity_intros) show "is_pole (\w. f w / g w) z" proof (rule is_pole_divide) have "\\<^sub>F x in at z. g x \ 0" apply (rule non_zero_neighbour) using g_nconst by auto moreover have "g \z\ 0" using DERIV_isCont assms(8) continuous_at g_deriv by force ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp show "isCont f z" using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on by auto show "f z \ 0" by fact qed show "filterlim id (at z) (at z)" by (simp add: filterlim_iff) have "((\w. (f w * (w - z)) / g w) \ f z / g') (at z)" proof (rule lhopital_complex_simple) show "((\w. f w * (w - z)) has_field_derivative f z) (at z)" using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo]) show "(g has_field_derivative g') (at z)" by fact qed (insert assms, auto) then show "((\w. (f w / g w) * (w - z)) \ f z / g') (at z)" by (simp add: field_split_simps) qed qed subsection \The argument principle\ theorem argument_principle: fixes f::"complex \ complex" and poles s:: "complex set" defines "pz \ {w. f w = 0 \ w \ poles}" \ \\<^term>\pz\ is the set of poles and zeros\ assumes "open s" and "connected s" and f_holo:"f holomorphic_on s-poles" and h_holo:"h holomorphic_on s" and "valid_path g" and loop:"pathfinish g = pathstart g" and path_img:"path_image g \ s - pz" and homo:"\z. (z \ s) \ winding_number g z = 0" and finite:"finite pz" and poles:"\p\poles. is_pole f p" shows "contour_integral g (\x. deriv f x * h x / f x) = 2 * pi * \ * (\p\pz. winding_number g p * h p * zorder f p)" (is "?L=?R") proof - define c where "c \ 2 * complex_of_real pi * \ " define ff where "ff \ (\x. deriv f x * h x / f x)" define cont where "cont \ \ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" define avoid where "avoid \ \p e. \w\cball p e. w \ s \ (w \ p \ w \ pz)" have "\e>0. avoid p e \ (p\pz \ cont ff p e)" when "p\s" for p proof - obtain e1 where "e1>0" and e1_avoid:"avoid p e1" using finite_cball_avoid[OF \open s\ finite] \p\s\ unfolding avoid_def by auto have "\e2>0. cball p e2 \ ball p e1 \ cont ff p e2" when "p\pz" proof - define po where "po \ zorder f p" define pp where "pp \ zor_poly f p" define f' where "f' \ \w. pp w * (w - p) powr po" define ff' where "ff' \ (\x. deriv f' x * h x / f' x)" obtain r where "pp p\0" "r>0" and "rw\cball p r-{p}. f w = pp w * (w - p) powr po \ pp w \ 0)" proof - have "isolated_singularity_at f p" proof - have "f holomorphic_on ball p e1 - {p}" apply (intro holomorphic_on_subset[OF f_holo]) using e1_avoid \p\pz\ unfolding avoid_def pz_def by force then show ?thesis unfolding isolated_singularity_at_def using \e1>0\ analytic_on_open open_delete by blast qed moreover have "not_essential f p" proof (cases "is_pole f p") case True then show ?thesis unfolding not_essential_def by auto next case False then have "p\s-poles" using \p\s\ poles unfolding pz_def by auto moreover have "open (s-poles)" using \open s\ apply (elim open_Diff) apply (rule finite_imp_closed) using finite unfolding pz_def by simp ultimately have "isCont f p" using holomorphic_on_imp_continuous_on[OF f_holo] continuous_on_eq_continuous_at by auto then show ?thesis unfolding isCont_def not_essential_def by auto qed moreover have "\\<^sub>F w in at p. f w \ 0 " proof (rule ccontr) assume "\ (\\<^sub>F w in at p. f w \ 0)" then have "\\<^sub>F w in at p. f w= 0" unfolding frequently_def by auto then obtain rr where "rr>0" "\w\ball p rr - {p}. f w =0" unfolding eventually_at by (auto simp add:dist_commute) then have "ball p rr - {p} \ {w\ball p rr-{p}. f w=0}" by blast moreover have "infinite (ball p rr - {p})" using \rr>0\ using finite_imp_not_open by fastforce ultimately have "infinite {w\ball p rr-{p}. f w=0}" using infinite_super by blast then have "infinite pz" unfolding pz_def infinite_super by auto then show False using \finite pz\ by auto qed ultimately obtain r where "pp p \ 0" and r:"r>0" "pp holomorphic_on cball p r" "(\w\cball p r - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" using zorder_exist[of f p,folded po_def pp_def] by auto define r1 where "r1=min r e1 / 2" have "r1e1>0\ \r>0\ by auto moreover have "r1>0" "pp holomorphic_on cball p r1" "(\w\cball p r1 - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" unfolding r1_def using \e1>0\ r by auto ultimately show ?thesis using that \pp p\0\ by auto qed define e2 where "e2 \ r/2" have "e2>0" using \r>0\ unfolding e2_def by auto define anal where "anal \ \w. deriv pp w * h w / pp w" define prin where "prin \ \w. po * h w / (w - p)" have "((\w. prin w + anal w) has_contour_integral c * po * h p) (circlepath p e2)" proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) have "ball p r \ s" using \r avoid_def ball_subset_cball e1_avoid by (simp add: subset_eq) then have "cball p e2 \ s" using \r>0\ unfolding e2_def by auto then have "(\w. po * h w) holomorphic_on cball p e2" using h_holo by (auto intro!: holomorphic_intros) then show "(prin has_contour_integral c * po * h p ) (circlepath p e2)" using Cauchy_integral_circlepath_simple[folded c_def, of "\w. po * h w"] \e2>0\ unfolding prin_def by (auto simp add: mult.assoc) have "anal holomorphic_on ball p r" unfolding anal_def using pp_holo h_holo pp_po \ball p r \ s\ \pp p\0\ by (auto intro!: holomorphic_intros) then show "(anal has_contour_integral 0) (circlepath p e2)" using e2_def \r>0\ by (auto elim!: Cauchy_theorem_disc_simple) qed then have "cont ff' p e2" unfolding cont_def po_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto define wp where "wp \ w-p" have "wp\0" and "pp w \0" unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto moreover have der_f':"deriv f' w = po * pp w * (w-p) powr (po - 1) + deriv pp w * (w-p) powr po" proof (rule DERIV_imp_deriv) have "(pp has_field_derivative (deriv pp w)) (at w)" using DERIV_deriv_iff_has_field_derivative pp_holo \w\p\ by (meson open_ball \w \ ball p r\ ball_subset_cball holomorphic_derivI holomorphic_on_subset) then show " (f' has_field_derivative of_int po * pp w * (w - p) powr of_int (po - 1) + deriv pp w * (w - p) powr of_int po) (at w)" unfolding f'_def using \w\p\ by (auto intro!: derivative_eq_intros DERIV_cong[OF has_field_derivative_powr_of_int]) qed ultimately show "prin w + anal w = ff' w" unfolding ff'_def prin_def anal_def apply simp apply (unfold f'_def) apply (fold wp_def) apply (auto simp add:field_simps) by (metis (no_types, lifting) diff_add_cancel mult.commute powr_add powr_to_1) qed then have "cont ff p e2" unfolding cont_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto have "deriv f' w = deriv f w" proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo by (auto intro!: holomorphic_intros) next have "ball p e1 - {p} \ s - poles" using ball_subset_cball e1_avoid[unfolded avoid_def] unfolding pz_def by auto then have "ball p r - {p} \ s - poles" apply (elim dual_order.trans) using \r by auto then show "f holomorphic_on ball p r - {p}" using f_holo by auto next show "open (ball p r - {p})" by auto show "w \ ball p r - {p}" using \w\ball p r\ \w\p\ by auto next fix x assume "x \ ball p r - {p}" then show "f' x = f x" using pp_po unfolding f'_def by auto qed moreover have " f' w = f w " using \w \ ball p r\ ball_subset_cball subset_iff pp_po \w\p\ unfolding f'_def by auto ultimately show "ff' w = ff w" unfolding ff'_def ff_def by simp qed moreover have "cball p e2 \ ball p e1" using \0 < r\ \r e2_def by auto ultimately show ?thesis using \e2>0\ by auto qed then obtain e2 where e2:"p\pz \ e2>0 \ cball p e2 \ ball p e1 \ cont ff p e2" by auto define e4 where "e4 \ if p\pz then e2 else e1" have "e4>0" using e2 \e1>0\ unfolding e4_def by auto moreover have "avoid p e4" using e2 \e1>0\ e1_avoid unfolding e4_def avoid_def by auto moreover have "p\pz \ cont ff p e4" by (auto simp add: e2 e4_def) ultimately show ?thesis by auto qed then obtain get_e where get_e:"\p\s. get_e p>0 \ avoid p (get_e p) \ (p\pz \ cont ff p (get_e p))" by metis define ci where "ci \ \p. contour_integral (circlepath p (get_e p)) ff" define w where "w \ \p. winding_number g p" have "contour_integral g ff = (\p\pz. w p * ci p)" unfolding ci_def w_def proof (rule Cauchy_theorem_singularities[OF \open s\ \connected s\ finite _ \valid_path g\ loop path_img homo]) have "open (s - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto then show "ff holomorphic_on s - pz" unfolding ff_def using f_holo h_holo by (auto intro!: holomorphic_intros simp add:pz_def) next show "\p\s. 0 < get_e p \ (\w\cball p (get_e p). w \ s \ (w \ p \ w \ pz))" using get_e using avoid_def by blast qed also have "... = (\p\pz. c * w p * h p * zorder f p)" proof (rule sum.cong[of pz pz,simplified]) fix p assume "p \ pz" show "w p * ci p = c * w p * h p * (zorder f p)" proof (cases "p\s") assume "p \ s" have "ci p = c * h p * (zorder f p)" unfolding ci_def apply (rule contour_integral_unique) using get_e \p\s\ \p\pz\ unfolding cont_def by (metis mult.assoc mult.commute) thus ?thesis by auto next assume "p\s" then have "w p=0" using homo unfolding w_def by auto then show ?thesis by auto qed qed also have "... = c*(\p\pz. w p * h p * zorder f p)" unfolding sum_distrib_left by (simp add:algebra_simps) finally have "contour_integral g ff = c * (\p\pz. w p * h p * of_int (zorder f p))" . then show ?thesis unfolding ff_def c_def w_def by simp qed subsection \Rouche's theorem \ theorem Rouche_theorem: fixes f g::"complex \ complex" and s:: "complex set" defines "fg\(\p. f p + g p)" defines "zeros_fg\{p. fg p = 0}" and "zeros_f\{p. f p = 0}" assumes "open s" and "connected s" and "finite zeros_fg" and "finite zeros_f" and f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" and "valid_path \" and loop:"pathfinish \ = pathstart \" and path_img:"path_image \ \ s " and path_less:"\z\path_image \. cmod(f z) > cmod(g z)" and homo:"\z. (z \ s) \ winding_number \ z = 0" shows "(\p\zeros_fg. winding_number \ p * zorder fg p) = (\p\zeros_f. winding_number \ p * zorder f p)" proof - have path_fg:"path_image \ \ s - zeros_fg" proof - have False when "z\path_image \" and "f z + g z=0" for z proof - have "cmod (f z) > cmod (g z)" using \z\path_image \\ path_less by auto moreover have "f z = - g z" using \f z + g z =0\ by (simp add: eq_neg_iff_add_eq_0) then have "cmod (f z) = cmod (g z)" by auto ultimately show False by auto qed then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto qed have path_f:"path_image \ \ s - zeros_f" proof - have False when "z\path_image \" and "f z =0" for z proof - have "cmod (g z) < cmod (f z) " using \z\path_image \\ path_less by auto then have "cmod (g z) < 0" using \f z=0\ by auto then show False by auto qed then show ?thesis unfolding zeros_f_def using path_img by auto qed define w where "w \ \p. winding_number \ p" define c where "c \ 2 * complex_of_real pi * \" define h where "h \ \p. g p / f p + 1" obtain spikes where "finite spikes" and spikes: "\x\{0..1} - spikes. \ differentiable at x" using \valid_path \\ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have h_contour:"((\x. deriv h x / h x) has_contour_integral 0) \" proof - have outside_img:"0 \ outside (path_image (h o \))" proof - have "h p \ ball 1 1" when "p\path_image \" for p proof - have "cmod (g p/f p) <1" using path_less[rule_format,OF that] apply (cases "cmod (f p) = 0") by (auto simp add: norm_divide) then show ?thesis unfolding h_def by (auto simp add:dist_complex_def) qed then have "path_image (h o \) \ ball 1 1" by (simp add: image_subset_iff path_image_compose) moreover have " (0::complex) \ ball 1 1" by (simp add: dist_norm) ultimately show "?thesis" using convex_in_outside[of "ball 1 1" 0] outside_mono by blast qed have valid_h:"valid_path (h \ \)" proof (rule valid_path_compose_holomorphic[OF \valid_path \\ _ _ path_f]) show "h holomorphic_on s - zeros_f" unfolding h_def using f_holo g_holo by (auto intro!: holomorphic_intros simp add:zeros_f_def) next show "open (s - zeros_f)" using \finite zeros_f\ \open s\ finite_imp_closed by auto qed have "((\z. 1/z) has_contour_integral 0) (h \ \)" proof - have "0 \ path_image (h \ \)" using outside_img by (simp add: outside_def) then have "((\z. 1/z) has_contour_integral c * winding_number (h \ \) 0) (h \ \)" using has_contour_integral_winding_number[of "h o \" 0,simplified] valid_h unfolding c_def by auto moreover have "winding_number (h o \) 0 = 0" proof - have "0 \ outside (path_image (h \ \))" using outside_img . moreover have "path (h o \)" using valid_h by (simp add: valid_path_imp_path) moreover have "pathfinish (h o \) = pathstart (h o \)" by (simp add: loop pathfinish_compose pathstart_compose) ultimately show ?thesis using winding_number_zero_in_outside by auto qed ultimately show ?thesis by auto qed moreover have "vector_derivative (h \ \) (at x) = vector_derivative \ (at x) * deriv h (\ x)" when "x\{0..1} - spikes" for x proof (rule vector_derivative_chain_at_general) show "\ differentiable at x" using that \valid_path \\ spikes by auto next define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" define t where "t \ \ x" have "f t\0" unfolding zeros_f_def t_def by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) moreover have "t\s" using contra_subsetD path_image_def path_fg t_def that by fastforce ultimately have "(h has_field_derivative der t) (at t)" unfolding h_def der_def using g_holo f_holo \open s\ by (auto intro!: holomorphic_derivI derivative_eq_intros) then show "h field_differentiable at (\ x)" unfolding t_def field_differentiable_def by blast qed then have " ((/) 1 has_contour_integral 0) (h \ \) = ((\x. deriv h x / h x) has_contour_integral 0) \" unfolding has_contour_integral apply (intro has_integral_spike_eq[OF negligible_finite, OF \finite spikes\]) by auto ultimately show ?thesis by auto qed then have "contour_integral \ (\x. deriv h x / h x) = 0" using contour_integral_unique by simp moreover have "contour_integral \ (\x. deriv fg x / fg x) = contour_integral \ (\x. deriv f x / f x) + contour_integral \ (\p. deriv h p / h p)" proof - have "(\p. deriv f p / f p) contour_integrable_on \" proof (rule contour_integrable_holomorphic_simple[OF _ _ \valid_path \\ path_f]) show "open (s - zeros_f)" using finite_imp_closed[OF \finite zeros_f\] \open s\ by auto then show "(\p. deriv f p / f p) holomorphic_on s - zeros_f" using f_holo by (auto intro!: holomorphic_intros simp add:zeros_f_def) qed moreover have "(\p. deriv h p / h p) contour_integrable_on \" using h_contour by (simp add: has_contour_integral_integrable) ultimately have "contour_integral \ (\x. deriv f x / f x + deriv h x / h x) = contour_integral \ (\p. deriv f p / f p) + contour_integral \ (\p. deriv h p / h p)" using contour_integral_add[of "(\p. deriv f p / f p)" \ "(\p. deriv h p / h p)" ] by auto moreover have "deriv fg p / fg p = deriv f p / f p + deriv h p / h p" when "p\ path_image \" for p proof - have "fg p\0" and "f p\0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def by auto have "h p\0" proof (rule ccontr) assume "\ h p \ 0" then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) then have "cmod (g p/f p) = 1" by auto moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] apply (cases "cmod (f p) = 0") by (auto simp add: norm_divide) ultimately show False by auto qed have der_fg:"deriv fg p = deriv f p + deriv g p" unfolding fg_def using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _ \open s\] path_img that by auto have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" proof - define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" have "p\s" using path_img that by auto then have "(h has_field_derivative der p) (at p)" unfolding h_def der_def using g_holo f_holo \open s\ \f p\0\ by (auto intro!: derivative_eq_intros holomorphic_derivI) then show ?thesis unfolding der_def using DERIV_imp_deriv by auto qed show ?thesis apply (simp only:der_fg der_h) apply (auto simp add:field_simps \h p\0\ \f p\0\ \fg p\0\) by (auto simp add:field_simps h_def \f p\0\ fg_def) qed then have "contour_integral \ (\p. deriv fg p / fg p) = contour_integral \ (\p. deriv f p / f p + deriv h p / h p)" by (elim contour_integral_eq) ultimately show ?thesis by auto qed moreover have "contour_integral \ (\x. deriv fg x / fg x) = c * (\p\zeros_fg. w p * zorder fg p)" unfolding c_def zeros_fg_def w_def proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo , of _ "{}" "\_. 1",simplified]) show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto show "path_image \ \ s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def . show " finite {p. fg p = 0}" using \finite zeros_fg\ unfolding zeros_fg_def . qed moreover have "contour_integral \ (\x. deriv f x / f x) = c * (\p\zeros_f. w p * zorder f p)" unfolding c_def zeros_f_def w_def proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo , of _ "{}" "\_. 1",simplified]) show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto show "path_image \ \ s - {p. f p = 0}" using path_f unfolding zeros_f_def . show " finite {p. f p = 0}" using \finite zeros_f\ unfolding zeros_f_def . qed ultimately have " c* (\p\zeros_fg. w p * (zorder fg p)) = c* (\p\zeros_f. w p * (zorder f p))" by auto then show ?thesis unfolding c_def using w_def by auto qed end