src/HOL/Analysis/Conformal_Mappings.thy
changeset 71189 954ee5acaae0
parent 71181 8331063570d6
child 71190 8b8f9d3b3fac
equal deleted inserted replaced
71181:8331063570d6 71189:954ee5acaae0
     1 section \<open>Conformal Mappings and Consequences of Cauchy's Integral Theorem\<close>
       
     2 
       
     3 text\<open>By John Harrison et al.  Ported from HOL Light by L C Paulson (2016)\<close>
       
     4 
       
     5 text\<open>Also Cauchy's residue theorem by Wenda Li (2016)\<close>
       
     6 
       
     7 theory Conformal_Mappings
       
     8 imports Cauchy_Integral_Theorem
       
     9 
       
    10 begin
       
    11 
       
    12 (* FIXME mv to Cauchy_Integral_Theorem.thy *)
       
    13 subsection\<open>Cauchy's inequality and more versions of Liouville\<close>
       
    14 
       
    15 lemma Cauchy_higher_deriv_bound:
       
    16     assumes holf: "f holomorphic_on (ball z r)"
       
    17         and contf: "continuous_on (cball z r) f"
       
    18         and fin : "\<And>w. w \<in> ball z r \<Longrightarrow> f w \<in> ball y B0"
       
    19         and "0 < r" and "0 < n"
       
    20       shows "norm ((deriv ^^ n) f z) \<le> (fact n) * B0 / r^n"
       
    21 proof -
       
    22   have "0 < B0" using \<open>0 < r\<close> fin [of z]
       
    23     by (metis ball_eq_empty ex_in_conv fin not_less)
       
    24   have le_B0: "\<And>w. cmod (w - z) \<le> r \<Longrightarrow> cmod (f w - y) \<le> B0"
       
    25     apply (rule continuous_on_closure_norm_le [of "ball z r" "\<lambda>w. f w - y"])
       
    26     apply (auto simp: \<open>0 < r\<close>  dist_norm norm_minus_commute)
       
    27     apply (rule continuous_intros contf)+
       
    28     using fin apply (simp add: dist_commute dist_norm less_eq_real_def)
       
    29     done
       
    30   have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w) z - (deriv ^^ n) (\<lambda>w. y) z"
       
    31     using \<open>0 < n\<close> by simp
       
    32   also have "... = (deriv ^^ n) (\<lambda>w. f w - y) z"
       
    33     by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \<open>0 < r\<close>)
       
    34   finally have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w - y) z" .
       
    35   have contf': "continuous_on (cball z r) (\<lambda>u. f u - y)"
       
    36     by (rule contf continuous_intros)+
       
    37   have holf': "(\<lambda>u. (f u - y)) holomorphic_on (ball z r)"
       
    38     by (simp add: holf holomorphic_on_diff)
       
    39   define a where "a = (2 * pi)/(fact n)"
       
    40   have "0 < a"  by (simp add: a_def)
       
    41   have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)"
       
    42     using \<open>0 < r\<close> by (simp add: a_def field_split_simps)
       
    43   have der_dif: "(deriv ^^ n) (\<lambda>w. f w - y) z = (deriv ^^ n) f z"
       
    44     using \<open>0 < r\<close> \<open>0 < n\<close>
       
    45     by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const])
       
    46   have "norm ((2 * of_real pi * \<i>)/(fact n) * (deriv ^^ n) (\<lambda>w. f w - y) z)
       
    47         \<le> (B0/r^(Suc n)) * (2 * pi * r)"
       
    48     apply (rule has_contour_integral_bound_circlepath [of "(\<lambda>u. (f u - y)/(u - z)^(Suc n))" _ z])
       
    49     using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf']
       
    50     using \<open>0 < B0\<close> \<open>0 < r\<close>
       
    51     apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0)
       
    52     done
       
    53   then show ?thesis
       
    54     using \<open>0 < r\<close>
       
    55     by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0)
       
    56 qed
       
    57 
       
    58 lemma Cauchy_inequality:
       
    59     assumes holf: "f holomorphic_on (ball \<xi> r)"
       
    60         and contf: "continuous_on (cball \<xi> r) f"
       
    61         and "0 < r"
       
    62         and nof: "\<And>x. norm(\<xi>-x) = r \<Longrightarrow> norm(f x) \<le> B"
       
    63       shows "norm ((deriv ^^ n) f \<xi>) \<le> (fact n) * B / r^n"
       
    64 proof -
       
    65   obtain x where "norm (\<xi>-x) = r"
       
    66     by (metis abs_of_nonneg add_diff_cancel_left' \<open>0 < r\<close> diff_add_cancel
       
    67                  dual_order.strict_implies_order norm_of_real)
       
    68   then have "0 \<le> B"
       
    69     by (metis nof norm_not_less_zero not_le order_trans)
       
    70   have  "((\<lambda>u. f u / (u - \<xi>) ^ Suc n) has_contour_integral (2 * pi) * \<i> / fact n * (deriv ^^ n) f \<xi>)
       
    71          (circlepath \<xi> r)"
       
    72     apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf])
       
    73     using \<open>0 < r\<close> by simp
       
    74   then have "norm ((2 * pi * \<i>)/(fact n) * (deriv ^^ n) f \<xi>) \<le> (B / r^(Suc n)) * (2 * pi * r)"
       
    75     apply (rule has_contour_integral_bound_circlepath)
       
    76     using \<open>0 \<le> B\<close> \<open>0 < r\<close>
       
    77     apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc)
       
    78     done
       
    79   then show ?thesis using \<open>0 < r\<close>
       
    80     by (simp add: norm_divide norm_mult field_simps)
       
    81 qed
       
    82 
       
    83 lemma Liouville_polynomial:
       
    84     assumes holf: "f holomorphic_on UNIV"
       
    85         and nof: "\<And>z. A \<le> norm z \<Longrightarrow> norm(f z) \<le> B * norm z ^ n"
       
    86       shows "f \<xi> = (\<Sum>k\<le>n. (deriv^^k) f 0 / fact k * \<xi> ^ k)"
       
    87 proof (cases rule: le_less_linear [THEN disjE])
       
    88   assume "B \<le> 0"
       
    89   then have "\<And>z. A \<le> norm z \<Longrightarrow> norm(f z) = 0"
       
    90     by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le)
       
    91   then have f0: "(f \<longlongrightarrow> 0) at_infinity"
       
    92     using Lim_at_infinity by force
       
    93   then have [simp]: "f = (\<lambda>w. 0)"
       
    94     using Liouville_weak [OF holf, of 0]
       
    95     by (simp add: eventually_at_infinity f0) meson
       
    96   show ?thesis by simp
       
    97 next
       
    98   assume "0 < B"
       
    99   have "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * (\<xi> - 0)^k) sums f \<xi>)"
       
   100     apply (rule holomorphic_power_series [where r = "norm \<xi> + 1"])
       
   101     using holf holomorphic_on_subset apply auto
       
   102     done
       
   103   then have sumsf: "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * \<xi>^k) sums f \<xi>)" by simp
       
   104   have "(deriv ^^ k) f 0 / fact k * \<xi> ^ k = 0" if "k>n" for k
       
   105   proof (cases "(deriv ^^ k) f 0 = 0")
       
   106     case True then show ?thesis by simp
       
   107   next
       
   108     case False
       
   109     define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
       
   110     have "1 \<le> abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
       
   111       using \<open>0 < B\<close> by simp
       
   112     then have wge1: "1 \<le> norm w"
       
   113       by (metis norm_of_real w_def)
       
   114     then have "w \<noteq> 0" by auto
       
   115     have kB: "0 < fact k * B"
       
   116       using \<open>0 < B\<close> by simp
       
   117     then have "0 \<le> fact k * B / cmod ((deriv ^^ k) f 0)"
       
   118       by simp
       
   119     then have wgeA: "A \<le> cmod w"
       
   120       by (simp only: w_def norm_of_real)
       
   121     have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
       
   122       using \<open>0 < B\<close> by simp
       
   123     then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w"
       
   124       by (metis norm_of_real w_def)
       
   125     then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)"
       
   126       using False by (simp add: field_split_simps mult.commute split: if_split_asm)
       
   127     also have "... \<le> fact k * (B * norm w ^ n) / norm w ^ k"
       
   128       apply (rule Cauchy_inequality)
       
   129          using holf holomorphic_on_subset apply force
       
   130         using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast
       
   131        using \<open>w \<noteq> 0\<close> apply simp
       
   132        by (metis nof wgeA dist_0_norm dist_norm)
       
   133     also have "... = fact k * (B * 1 / cmod w ^ (k-n))"
       
   134       apply (simp only: mult_cancel_left times_divide_eq_right [symmetric])
       
   135       using \<open>k>n\<close> \<open>w \<noteq> 0\<close> \<open>0 < B\<close> apply (simp add: field_split_simps semiring_normalization_rules)
       
   136       done
       
   137     also have "... = fact k * B / cmod w ^ (k-n)"
       
   138       by simp
       
   139     finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" .
       
   140     then have "1 / cmod w < 1 / cmod w ^ (k - n)"
       
   141       by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos)
       
   142     then have "cmod w ^ (k - n) < cmod w"
       
   143       by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one)
       
   144     with self_le_power [OF wge1] have False
       
   145       by (meson diff_is_0_eq not_gr0 not_le that)
       
   146     then show ?thesis by blast
       
   147   qed
       
   148   then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \<xi> ^ (k + Suc n) = 0" for k
       
   149     using not_less_eq by blast
       
   150   then have "(\<lambda>i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \<xi> ^ (i + Suc n)) sums 0"
       
   151     by (rule sums_0)
       
   152   with sums_split_initial_segment [OF sumsf, where n = "Suc n"]
       
   153   show ?thesis
       
   154     using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce
       
   155 qed
       
   156 
       
   157 text\<open>Every bounded entire function is a constant function.\<close>
       
   158 theorem Liouville_theorem:
       
   159     assumes holf: "f holomorphic_on UNIV"
       
   160         and bf: "bounded (range f)"
       
   161     obtains c where "\<And>z. f z = c"
       
   162 proof -
       
   163   obtain B where "\<And>z. cmod (f z) \<le> B"
       
   164     by (meson bf bounded_pos rangeI)
       
   165   then show ?thesis
       
   166     using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast
       
   167 qed
       
   168 
       
   169 text\<open>A holomorphic function f has only isolated zeros unless f is 0.\<close>
       
   170 
       
   171 lemma powser_0_nonzero:
       
   172   fixes a :: "nat \<Rightarrow> 'a::{real_normed_field,banach}"
       
   173   assumes r: "0 < r"
       
   174       and sm: "\<And>x. norm (x - \<xi>) < r \<Longrightarrow> (\<lambda>n. a n * (x - \<xi>) ^ n) sums (f x)"
       
   175       and [simp]: "f \<xi> = 0"
       
   176       and m0: "a m \<noteq> 0" and "m>0"
       
   177   obtains s where "0 < s" and "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0"
       
   178 proof -
       
   179   have "r \<le> conv_radius a"
       
   180     using sm sums_summable by (auto simp: le_conv_radius_iff [where \<xi>=\<xi>])
       
   181   obtain m where am: "a m \<noteq> 0" and az [simp]: "(\<And>n. n<m \<Longrightarrow> a n = 0)"
       
   182     apply (rule_tac m = "LEAST n. a n \<noteq> 0" in that)
       
   183     using m0
       
   184     apply (rule LeastI2)
       
   185     apply (fastforce intro:  dest!: not_less_Least)+
       
   186     done
       
   187   define b where "b i = a (i+m) / a m" for i
       
   188   define g where "g x = suminf (\<lambda>i. b i * (x - \<xi>) ^ i)" for x
       
   189   have [simp]: "b 0 = 1"
       
   190     by (simp add: am b_def)
       
   191   { fix x::'a
       
   192     assume "norm (x - \<xi>) < r"
       
   193     then have "(\<lambda>n. (a m * (x - \<xi>)^m) * (b n * (x - \<xi>)^n)) sums (f x)"
       
   194       using am az sm sums_zero_iff_shift [of m "(\<lambda>n. a n * (x - \<xi>) ^ n)" "f x"]
       
   195       by (simp add: b_def monoid_mult_class.power_add algebra_simps)
       
   196     then have "x \<noteq> \<xi> \<Longrightarrow> (\<lambda>n. b n * (x - \<xi>)^n) sums (f x / (a m * (x - \<xi>)^m))"
       
   197       using am by (simp add: sums_mult_D)
       
   198   } note bsums = this
       
   199   then have  "norm (x - \<xi>) < r \<Longrightarrow> summable (\<lambda>n. b n * (x - \<xi>)^n)" for x
       
   200     using sums_summable by (cases "x=\<xi>") auto
       
   201   then have "r \<le> conv_radius b"
       
   202     by (simp add: le_conv_radius_iff [where \<xi>=\<xi>])
       
   203   then have "r/2 < conv_radius b"
       
   204     using not_le order_trans r by fastforce
       
   205   then have "continuous_on (cball \<xi> (r/2)) g"
       
   206     using powser_continuous_suminf [of "r/2" b \<xi>] by (simp add: g_def)
       
   207   then obtain s where "s>0"  "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> dist (g x) (g \<xi>) < 1/2"
       
   208     apply (rule continuous_onE [where x=\<xi> and e = "1/2"])
       
   209     using r apply (auto simp: norm_minus_commute dist_norm)
       
   210     done
       
   211   moreover have "g \<xi> = 1"
       
   212     by (simp add: g_def)
       
   213   ultimately have gnz: "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> (g x) \<noteq> 0"
       
   214     by fastforce
       
   215   have "f x \<noteq> 0" if "x \<noteq> \<xi>" "norm (x - \<xi>) \<le> s" "norm (x - \<xi>) \<le> r/2" for x
       
   216     using bsums [of x] that gnz [of x]
       
   217     apply (auto simp: g_def)
       
   218     using r sums_iff by fastforce
       
   219   then show ?thesis
       
   220     apply (rule_tac s="min s (r/2)" in that)
       
   221     using \<open>0 < r\<close> \<open>0 < s\<close> by (auto simp: dist_commute dist_norm)
       
   222 qed
       
   223 
       
   224 subsection \<open>Analytic continuation\<close>
       
   225 
       
   226 proposition isolated_zeros:
       
   227   assumes holf: "f holomorphic_on S"
       
   228       and "open S" "connected S" "\<xi> \<in> S" "f \<xi> = 0" "\<beta> \<in> S" "f \<beta> \<noteq> 0"
       
   229     obtains r where "0 < r" and "ball \<xi> r \<subseteq> S" and
       
   230         "\<And>z. z \<in> ball \<xi> r - {\<xi>} \<Longrightarrow> f z \<noteq> 0"
       
   231 proof -
       
   232   obtain r where "0 < r" and r: "ball \<xi> r \<subseteq> S"
       
   233     using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_ball_eq by blast
       
   234   have powf: "((\<lambda>n. (deriv ^^ n) f \<xi> / (fact n) * (z - \<xi>)^n) sums f z)" if "z \<in> ball \<xi> r" for z
       
   235     apply (rule holomorphic_power_series [OF _ that])
       
   236     apply (rule holomorphic_on_subset [OF holf r])
       
   237     done
       
   238   obtain m where m: "(deriv ^^ m) f \<xi> / (fact m) \<noteq> 0"
       
   239     using holomorphic_fun_eq_0_on_connected [OF holf \<open>open S\<close> \<open>connected S\<close> _ \<open>\<xi> \<in> S\<close> \<open>\<beta> \<in> S\<close>] \<open>f \<beta> \<noteq> 0\<close>
       
   240     by auto
       
   241   then have "m \<noteq> 0" using assms(5) funpow_0 by fastforce
       
   242   obtain s where "0 < s" and s: "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0"
       
   243     apply (rule powser_0_nonzero [OF \<open>0 < r\<close> powf \<open>f \<xi> = 0\<close> m])
       
   244     using \<open>m \<noteq> 0\<close> by (auto simp: dist_commute dist_norm)
       
   245   have "0 < min r s"  by (simp add: \<open>0 < r\<close> \<open>0 < s\<close>)
       
   246   then show ?thesis
       
   247     apply (rule that)
       
   248     using r s by auto
       
   249 qed
       
   250 
       
   251 proposition analytic_continuation:
       
   252   assumes holf: "f holomorphic_on S"
       
   253       and "open S" and "connected S"
       
   254       and "U \<subseteq> S" and "\<xi> \<in> S"
       
   255       and "\<xi> islimpt U"
       
   256       and fU0 [simp]: "\<And>z. z \<in> U \<Longrightarrow> f z = 0"
       
   257       and "w \<in> S"
       
   258     shows "f w = 0"
       
   259 proof -
       
   260   obtain e where "0 < e" and e: "cball \<xi> e \<subseteq> S"
       
   261     using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_cball_eq by blast
       
   262   define T where "T = cball \<xi> e \<inter> U"
       
   263   have contf: "continuous_on (closure T) f"
       
   264     by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on
       
   265               holomorphic_on_subset inf.cobounded1)
       
   266   have fT0 [simp]: "\<And>x. x \<in> T \<Longrightarrow> f x = 0"
       
   267     by (simp add: T_def)
       
   268   have "\<And>r. \<lbrakk>\<forall>e>0. \<exists>x'\<in>U. x' \<noteq> \<xi> \<and> dist x' \<xi> < e; 0 < r\<rbrakk> \<Longrightarrow> \<exists>x'\<in>cball \<xi> e \<inter> U. x' \<noteq> \<xi> \<and> dist x' \<xi> < r"
       
   269     by (metis \<open>0 < e\<close> IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj)
       
   270   then have "\<xi> islimpt T" using \<open>\<xi> islimpt U\<close>
       
   271     by (auto simp: T_def islimpt_approachable)
       
   272   then have "\<xi> \<in> closure T"
       
   273     by (simp add: closure_def)
       
   274   then have "f \<xi> = 0"
       
   275     by (auto simp: continuous_constant_on_closure [OF contf])
       
   276   show ?thesis
       
   277     apply (rule ccontr)
       
   278     apply (rule isolated_zeros [OF holf \<open>open S\<close> \<open>connected S\<close> \<open>\<xi> \<in> S\<close> \<open>f \<xi> = 0\<close> \<open>w \<in> S\<close>], assumption)
       
   279     by (metis open_ball \<open>\<xi> islimpt T\<close> centre_in_ball fT0 insertE insert_Diff islimptE)
       
   280 qed
       
   281 
       
   282 corollary analytic_continuation_open:
       
   283   assumes "open s" and "open s'" and "s \<noteq> {}" and "connected s'"
       
   284       and "s \<subseteq> s'"
       
   285   assumes "f holomorphic_on s'" and "g holomorphic_on s'"
       
   286       and "\<And>z. z \<in> s \<Longrightarrow> f z = g z"
       
   287   assumes "z \<in> s'"
       
   288   shows   "f z = g z"
       
   289 proof -
       
   290   from \<open>s \<noteq> {}\<close> obtain \<xi> where "\<xi> \<in> s" by auto
       
   291   with \<open>open s\<close> have \<xi>: "\<xi> islimpt s"
       
   292     by (intro interior_limit_point) (auto simp: interior_open)
       
   293   have "f z - g z = 0"
       
   294     by (rule analytic_continuation[of "\<lambda>z. f z - g z" s' s \<xi>])
       
   295        (insert assms \<open>\<xi> \<in> s\<close> \<xi>, auto intro: holomorphic_intros)
       
   296   thus ?thesis by simp
       
   297 qed
       
   298 
       
   299 subsection\<open>Open mapping theorem\<close>
       
   300 
       
   301 lemma holomorphic_contract_to_zero:
       
   302   assumes contf: "continuous_on (cball \<xi> r) f"
       
   303       and holf: "f holomorphic_on ball \<xi> r"
       
   304       and "0 < r"
       
   305       and norm_less: "\<And>z. norm(\<xi> - z) = r \<Longrightarrow> norm(f \<xi>) < norm(f z)"
       
   306   obtains z where "z \<in> ball \<xi> r" "f z = 0"
       
   307 proof -
       
   308   { assume fnz: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w \<noteq> 0"
       
   309     then have "0 < norm (f \<xi>)"
       
   310       by (simp add: \<open>0 < r\<close>)
       
   311     have fnz': "\<And>w. w \<in> cball \<xi> r \<Longrightarrow> f w \<noteq> 0"
       
   312       by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero)
       
   313     have "frontier(cball \<xi> r) \<noteq> {}"
       
   314       using \<open>0 < r\<close> by simp
       
   315     define g where [abs_def]: "g z = inverse (f z)" for z
       
   316     have contg: "continuous_on (cball \<xi> r) g"
       
   317       unfolding g_def using contf continuous_on_inverse fnz' by blast
       
   318     have holg: "g holomorphic_on ball \<xi> r"
       
   319       unfolding g_def using fnz holf holomorphic_on_inverse by blast
       
   320     have "frontier (cball \<xi> r) \<subseteq> cball \<xi> r"
       
   321       by (simp add: subset_iff)
       
   322     then have contf': "continuous_on (frontier (cball \<xi> r)) f"
       
   323           and contg': "continuous_on (frontier (cball \<xi> r)) g"
       
   324       by (blast intro: contf contg continuous_on_subset)+
       
   325     have froc: "frontier(cball \<xi> r) \<noteq> {}"
       
   326       using \<open>0 < r\<close> by simp
       
   327     moreover have "continuous_on (frontier (cball \<xi> r)) (norm o f)"
       
   328       using contf' continuous_on_compose continuous_on_norm_id by blast
       
   329     ultimately obtain w where w: "w \<in> frontier(cball \<xi> r)"
       
   330                           and now: "\<And>x. x \<in> frontier(cball \<xi> r) \<Longrightarrow> norm (f w) \<le> norm (f x)"
       
   331       apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]])
       
   332       apply simp
       
   333       done
       
   334     then have fw: "0 < norm (f w)"
       
   335       by (simp add: fnz')
       
   336     have "continuous_on (frontier (cball \<xi> r)) (norm o g)"
       
   337       using contg' continuous_on_compose continuous_on_norm_id by blast
       
   338     then obtain v where v: "v \<in> frontier(cball \<xi> r)"
       
   339                and nov: "\<And>x. x \<in> frontier(cball \<xi> r) \<Longrightarrow> norm (g v) \<ge> norm (g x)"
       
   340       apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]])
       
   341       apply simp
       
   342       done
       
   343     then have fv: "0 < norm (f v)"
       
   344       by (simp add: fnz')
       
   345     have "norm ((deriv ^^ 0) g \<xi>) \<le> fact 0 * norm (g v) / r ^ 0"
       
   346       by (rule Cauchy_inequality [OF holg contg \<open>0 < r\<close>]) (simp add: dist_norm nov)
       
   347     then have "cmod (g \<xi>) \<le> norm (g v)"
       
   348       by simp
       
   349     with w have wr: "norm (\<xi> - w) = r" and nfw: "norm (f w) \<le> norm (f \<xi>)"
       
   350       apply (simp_all add: dist_norm)
       
   351       by (metis \<open>0 < cmod (f \<xi>)\<close> g_def less_imp_inverse_less norm_inverse not_le now order_trans v)
       
   352     with fw have False
       
   353       using norm_less by force
       
   354   }
       
   355   with that show ?thesis by blast
       
   356 qed
       
   357 
       
   358 theorem open_mapping_thm:
       
   359   assumes holf: "f holomorphic_on S"
       
   360       and S: "open S" and "connected S"
       
   361       and "open U" and "U \<subseteq> S"
       
   362       and fne: "\<not> f constant_on S"
       
   363     shows "open (f ` U)"
       
   364 proof -
       
   365   have *: "open (f ` U)"
       
   366           if "U \<noteq> {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\<And>x. \<exists>y \<in> U. f y \<noteq> x"
       
   367           for U
       
   368   proof (clarsimp simp: open_contains_ball)
       
   369     fix \<xi> assume \<xi>: "\<xi> \<in> U"
       
   370     show "\<exists>e>0. ball (f \<xi>) e \<subseteq> f ` U"
       
   371     proof -
       
   372       have hol: "(\<lambda>z. f z - f \<xi>) holomorphic_on U"
       
   373         by (rule holomorphic_intros that)+
       
   374       obtain s where "0 < s" and sbU: "ball \<xi> s \<subseteq> U"
       
   375                  and sne: "\<And>z. z \<in> ball \<xi> s - {\<xi>} \<Longrightarrow> (\<lambda>z. f z - f \<xi>) z \<noteq> 0"
       
   376         using isolated_zeros [OF hol U \<xi>]  by (metis fneU right_minus_eq)
       
   377       obtain r where "0 < r" and r: "cball \<xi> r \<subseteq> ball \<xi> s"
       
   378         apply (rule_tac r="s/2" in that)
       
   379         using \<open>0 < s\<close> by auto
       
   380       have "cball \<xi> r \<subseteq> U"
       
   381         using sbU r by blast
       
   382       then have frsbU: "frontier (cball \<xi> r) \<subseteq> U"
       
   383         using Diff_subset frontier_def order_trans by fastforce
       
   384       then have cof: "compact (frontier(cball \<xi> r))"
       
   385         by blast
       
   386       have frne: "frontier (cball \<xi> r) \<noteq> {}"
       
   387         using \<open>0 < r\<close> by auto
       
   388       have contfr: "continuous_on (frontier (cball \<xi> r)) (\<lambda>z. norm (f z - f \<xi>))"
       
   389         by (metis continuous_on_norm continuous_on_subset frsbU hol holomorphic_on_imp_continuous_on)
       
   390       obtain w where "norm (\<xi> - w) = r"
       
   391                  and w: "(\<And>z. norm (\<xi> - z) = r \<Longrightarrow> norm (f w - f \<xi>) \<le> norm(f z - f \<xi>))"
       
   392         apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]])
       
   393         apply (simp add: dist_norm)
       
   394         done
       
   395       moreover define \<epsilon> where "\<epsilon> \<equiv> norm (f w - f \<xi>) / 3"
       
   396       ultimately have "0 < \<epsilon>"
       
   397         using \<open>0 < r\<close> dist_complex_def r sne by auto
       
   398       have "ball (f \<xi>) \<epsilon> \<subseteq> f ` U"
       
   399       proof
       
   400         fix \<gamma>
       
   401         assume \<gamma>: "\<gamma> \<in> ball (f \<xi>) \<epsilon>"
       
   402         have *: "cmod (\<gamma> - f \<xi>) < cmod (\<gamma> - f z)" if "cmod (\<xi> - z) = r" for z
       
   403         proof -
       
   404           have lt: "cmod (f w - f \<xi>) / 3 < cmod (\<gamma> - f z)"
       
   405             using w [OF that] \<gamma>
       
   406             using dist_triangle2 [of "f \<xi>" "\<gamma>"  "f z"] dist_triangle2 [of "f \<xi>" "f z" \<gamma>]
       
   407             by (simp add: \<epsilon>_def dist_norm norm_minus_commute)
       
   408           show ?thesis
       
   409             by (metis \<epsilon>_def dist_commute dist_norm less_trans lt mem_ball \<gamma>)
       
   410        qed
       
   411        have "continuous_on (cball \<xi> r) (\<lambda>z. \<gamma> - f z)"
       
   412           apply (rule continuous_intros)+
       
   413           using \<open>cball \<xi> r \<subseteq> U\<close> \<open>f holomorphic_on U\<close>
       
   414           apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on)
       
   415           done
       
   416         moreover have "(\<lambda>z. \<gamma> - f z) holomorphic_on ball \<xi> r"
       
   417           apply (rule holomorphic_intros)+
       
   418           apply (metis \<open>cball \<xi> r \<subseteq> U\<close> \<open>f holomorphic_on U\<close> holomorphic_on_subset interior_cball interior_subset)
       
   419           done
       
   420         ultimately obtain z where "z \<in> ball \<xi> r" "\<gamma> - f z = 0"
       
   421           apply (rule holomorphic_contract_to_zero)
       
   422           apply (blast intro!: \<open>0 < r\<close> *)+
       
   423           done
       
   424         then show "\<gamma> \<in> f ` U"
       
   425           using \<open>cball \<xi> r \<subseteq> U\<close> by fastforce
       
   426       qed
       
   427       then show ?thesis using  \<open>0 < \<epsilon>\<close> by blast
       
   428     qed
       
   429   qed
       
   430   have "open (f ` X)" if "X \<in> components U" for X
       
   431   proof -
       
   432     have holfU: "f holomorphic_on U"
       
   433       using \<open>U \<subseteq> S\<close> holf holomorphic_on_subset by blast
       
   434     have "X \<noteq> {}"
       
   435       using that by (simp add: in_components_nonempty)
       
   436     moreover have "open X"
       
   437       using that \<open>open U\<close> open_components by auto
       
   438     moreover have "connected X"
       
   439       using that in_components_maximal by blast
       
   440     moreover have "f holomorphic_on X"
       
   441       by (meson that holfU holomorphic_on_subset in_components_maximal)
       
   442     moreover have "\<exists>y\<in>X. f y \<noteq> x" for x
       
   443     proof (rule ccontr)
       
   444       assume not: "\<not> (\<exists>y\<in>X. f y \<noteq> x)"
       
   445       have "X \<subseteq> S"
       
   446         using \<open>U \<subseteq> S\<close> in_components_subset that by blast
       
   447       obtain w where w: "w \<in> X" using \<open>X \<noteq> {}\<close> by blast
       
   448       have wis: "w islimpt X"
       
   449         using w \<open>open X\<close> interior_eq by auto
       
   450       have hol: "(\<lambda>z. f z - x) holomorphic_on S"
       
   451         by (simp add: holf holomorphic_on_diff)
       
   452       with fne [unfolded constant_on_def]
       
   453            analytic_continuation[OF hol S \<open>connected S\<close> \<open>X \<subseteq> S\<close> _ wis] not \<open>X \<subseteq> S\<close> w
       
   454       show False by auto
       
   455     qed
       
   456     ultimately show ?thesis
       
   457       by (rule *)
       
   458   qed
       
   459   then have "open (f ` \<Union>(components U))"
       
   460     by (metis (no_types, lifting) imageE image_Union open_Union)
       
   461   then show ?thesis
       
   462     by force
       
   463 qed
       
   464 
       
   465 text\<open>No need for \<^term>\<open>S\<close> to be connected. But the nonconstant condition is stronger.\<close>
       
   466 corollary\<^marker>\<open>tag unimportant\<close> open_mapping_thm2:
       
   467   assumes holf: "f holomorphic_on S"
       
   468       and S: "open S"
       
   469       and "open U" "U \<subseteq> S"
       
   470       and fnc: "\<And>X. \<lbrakk>open X; X \<subseteq> S; X \<noteq> {}\<rbrakk> \<Longrightarrow> \<not> f constant_on X"
       
   471     shows "open (f ` U)"
       
   472 proof -
       
   473   have "S = \<Union>(components S)" by simp
       
   474   with \<open>U \<subseteq> S\<close> have "U = (\<Union>C \<in> components S. C \<inter> U)" by auto
       
   475   then have "f ` U = (\<Union>C \<in> components S. f ` (C \<inter> U))"
       
   476     using image_UN by fastforce
       
   477   moreover
       
   478   { fix C assume "C \<in> components S"
       
   479     with S \<open>C \<in> components S\<close> open_components in_components_connected
       
   480     have C: "open C" "connected C" by auto
       
   481     have "C \<subseteq> S"
       
   482       by (metis \<open>C \<in> components S\<close> in_components_maximal)
       
   483     have nf: "\<not> f constant_on C"
       
   484       apply (rule fnc)
       
   485       using C \<open>C \<subseteq> S\<close> \<open>C \<in> components S\<close> in_components_nonempty by auto
       
   486     have "f holomorphic_on C"
       
   487       by (metis holf holomorphic_on_subset \<open>C \<subseteq> S\<close>)
       
   488     then have "open (f ` (C \<inter> U))"
       
   489       apply (rule open_mapping_thm [OF _ C _ _ nf])
       
   490       apply (simp add: C \<open>open U\<close> open_Int, blast)
       
   491       done
       
   492   } ultimately show ?thesis
       
   493     by force
       
   494 qed
       
   495 
       
   496 corollary\<^marker>\<open>tag unimportant\<close> open_mapping_thm3:
       
   497   assumes holf: "f holomorphic_on S"
       
   498       and "open S" and injf: "inj_on f S"
       
   499     shows  "open (f ` S)"
       
   500 apply (rule open_mapping_thm2 [OF holf])
       
   501 using assms
       
   502 apply (simp_all add:)
       
   503 using injective_not_constant subset_inj_on by blast
       
   504 
       
   505 subsection\<open>Maximum modulus principle\<close>
       
   506 
       
   507 text\<open>If \<^term>\<open>f\<close> is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is
       
   508    properly within the domain of \<^term>\<open>f\<close>.\<close>
       
   509 
       
   510 proposition maximum_modulus_principle:
       
   511   assumes holf: "f holomorphic_on S"
       
   512       and S: "open S" and "connected S"
       
   513       and "open U" and "U \<subseteq> S" and "\<xi> \<in> U"
       
   514       and no: "\<And>z. z \<in> U \<Longrightarrow> norm(f z) \<le> norm(f \<xi>)"
       
   515     shows "f constant_on S"
       
   516 proof (rule ccontr)
       
   517   assume "\<not> f constant_on S"
       
   518   then have "open (f ` U)"
       
   519     using open_mapping_thm assms by blast
       
   520   moreover have "\<not> open (f ` U)"
       
   521   proof -
       
   522     have "\<exists>t. cmod (f \<xi> - t) < e \<and> t \<notin> f ` U" if "0 < e" for e
       
   523       apply (rule_tac x="if 0 < Re(f \<xi>) then f \<xi> + (e/2) else f \<xi> - (e/2)" in exI)
       
   524       using that
       
   525       apply (simp add: dist_norm)
       
   526       apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym)
       
   527       done
       
   528     then show ?thesis
       
   529       unfolding open_contains_ball by (metis \<open>\<xi> \<in> U\<close> contra_subsetD dist_norm imageI mem_ball)
       
   530   qed
       
   531   ultimately show False
       
   532     by blast
       
   533 qed
       
   534 
       
   535 proposition maximum_modulus_frontier:
       
   536   assumes holf: "f holomorphic_on (interior S)"
       
   537       and contf: "continuous_on (closure S) f"
       
   538       and bos: "bounded S"
       
   539       and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> norm(f z) \<le> B"
       
   540       and "\<xi> \<in> S"
       
   541     shows "norm(f \<xi>) \<le> B"
       
   542 proof -
       
   543   have "compact (closure S)" using bos
       
   544     by (simp add: bounded_closure compact_eq_bounded_closed)
       
   545   moreover have "continuous_on (closure S) (cmod \<circ> f)"
       
   546     using contf continuous_on_compose continuous_on_norm_id by blast
       
   547   ultimately obtain z where zin: "z \<in> closure S" and z: "\<And>y. y \<in> closure S \<Longrightarrow> (cmod \<circ> f) y \<le> (cmod \<circ> f) z"
       
   548     using continuous_attains_sup [of "closure S" "norm o f"] \<open>\<xi> \<in> S\<close> by auto
       
   549   then consider "z \<in> frontier S" | "z \<in> interior S" using frontier_def by auto
       
   550   then have "norm(f z) \<le> B"
       
   551   proof cases
       
   552     case 1 then show ?thesis using leB by blast
       
   553   next
       
   554     case 2
       
   555     have zin: "z \<in> connected_component_set (interior S) z"
       
   556       by (simp add: 2)
       
   557     have "f constant_on (connected_component_set (interior S) z)"
       
   558       apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin])
       
   559       apply (metis connected_component_subset holf holomorphic_on_subset)
       
   560       apply (simp_all add: open_connected_component)
       
   561       by (metis closure_subset comp_eq_dest_lhs  interior_subset subsetCE z connected_component_in)
       
   562     then obtain c where c: "\<And>w. w \<in> connected_component_set (interior S) z \<Longrightarrow> f w = c"
       
   563       by (auto simp: constant_on_def)
       
   564     have "f ` closure(connected_component_set (interior S) z) \<subseteq> {c}"
       
   565       apply (rule image_closure_subset)
       
   566       apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset)
       
   567       using c
       
   568       apply auto
       
   569       done
       
   570     then have cc: "\<And>w. w \<in> closure(connected_component_set (interior S) z) \<Longrightarrow> f w = c" by blast
       
   571     have "frontier(connected_component_set (interior S) z) \<noteq> {}"
       
   572       apply (simp add: frontier_eq_empty)
       
   573       by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV)
       
   574     then obtain w where w: "w \<in> frontier(connected_component_set (interior S) z)"
       
   575        by auto
       
   576     then have "norm (f z) = norm (f w)"  by (simp add: "2" c cc frontier_def)
       
   577     also have "... \<le> B"
       
   578       apply (rule leB)
       
   579       using w
       
   580 using frontier_interior_subset frontier_of_connected_component_subset by blast
       
   581     finally show ?thesis .
       
   582   qed
       
   583   then show ?thesis
       
   584     using z \<open>\<xi> \<in> S\<close> closure_subset by fastforce
       
   585 qed
       
   586 
       
   587 corollary\<^marker>\<open>tag unimportant\<close> maximum_real_frontier:
       
   588   assumes holf: "f holomorphic_on (interior S)"
       
   589       and contf: "continuous_on (closure S) f"
       
   590       and bos: "bounded S"
       
   591       and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> Re(f z) \<le> B"
       
   592       and "\<xi> \<in> S"
       
   593     shows "Re(f \<xi>) \<le> B"
       
   594 using maximum_modulus_frontier [of "exp o f" S "exp B"]
       
   595       Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms
       
   596 by auto
       
   597 
       
   598 subsection\<^marker>\<open>tag unimportant\<close> \<open>Factoring out a zero according to its order\<close>
       
   599 
       
   600 lemma holomorphic_factor_order_of_zero:
       
   601   assumes holf: "f holomorphic_on S"
       
   602       and os: "open S"
       
   603       and "\<xi> \<in> S" "0 < n"
       
   604       and dnz: "(deriv ^^ n) f \<xi> \<noteq> 0"
       
   605       and dfz: "\<And>i. \<lbrakk>0 < i; i < n\<rbrakk> \<Longrightarrow> (deriv ^^ i) f \<xi> = 0"
       
   606    obtains g r where "0 < r"
       
   607                 "g holomorphic_on ball \<xi> r"
       
   608                 "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>)^n * g w"
       
   609                 "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
       
   610 proof -
       
   611   obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE)
       
   612   then have holfb: "f holomorphic_on ball \<xi> r"
       
   613     using holf holomorphic_on_subset by blast
       
   614   define g where "g w = suminf (\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i)" for w
       
   615   have sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w"
       
   616    and feq: "f w - f \<xi> = (w - \<xi>)^n * g w"
       
   617        if w: "w \<in> ball \<xi> r" for w
       
   618   proof -
       
   619     define powf where "powf = (\<lambda>i. (deriv ^^ i) f \<xi>/(fact i) * (w - \<xi>)^i)"
       
   620     have sing: "{..<n} - {i. powf i = 0} = (if f \<xi> = 0 then {} else {0})"
       
   621       unfolding powf_def using \<open>0 < n\<close> dfz by (auto simp: dfz; metis funpow_0 not_gr0)
       
   622     have "powf sums f w"
       
   623       unfolding powf_def by (rule holomorphic_power_series [OF holfb w])
       
   624     moreover have "(\<Sum>i<n. powf i) = f \<xi>"
       
   625       apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric])
       
   626       apply simp
       
   627       apply (simp only: dfz sing)
       
   628       apply (simp add: powf_def)
       
   629       done
       
   630     ultimately have fsums: "(\<lambda>i. powf (i+n)) sums (f w - f \<xi>)"
       
   631       using w sums_iff_shift' by metis
       
   632     then have *: "summable (\<lambda>i. (w - \<xi>) ^ n * ((deriv ^^ (i + n)) f \<xi> * (w - \<xi>) ^ i / fact (i + n)))"
       
   633       unfolding powf_def using sums_summable
       
   634       by (auto simp: power_add mult_ac)
       
   635     have "summable (\<lambda>i. (deriv ^^ (i + n)) f \<xi> * (w - \<xi>) ^ i / fact (i + n))"
       
   636     proof (cases "w=\<xi>")
       
   637       case False then show ?thesis
       
   638         using summable_mult [OF *, of "1 / (w - \<xi>) ^ n"] by simp
       
   639     next
       
   640       case True then show ?thesis
       
   641         by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"]
       
   642                  split: if_split_asm)
       
   643     qed
       
   644     then show sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w"
       
   645       by (simp add: summable_sums_iff g_def)
       
   646     show "f w - f \<xi> = (w - \<xi>)^n * g w"
       
   647       apply (rule sums_unique2)
       
   648       apply (rule fsums [unfolded powf_def])
       
   649       using sums_mult [OF sumsg, of "(w - \<xi>) ^ n"]
       
   650       by (auto simp: power_add mult_ac)
       
   651   qed
       
   652   then have holg: "g holomorphic_on ball \<xi> r"
       
   653     by (meson sumsg power_series_holomorphic)
       
   654   then have contg: "continuous_on (ball \<xi> r) g"
       
   655     by (blast intro: holomorphic_on_imp_continuous_on)
       
   656   have "g \<xi> \<noteq> 0"
       
   657     using dnz unfolding g_def
       
   658     by (subst suminf_finite [of "{0}"]) auto
       
   659   obtain d where "0 < d" and d: "\<And>w. w \<in> ball \<xi> d \<Longrightarrow> g w \<noteq> 0"
       
   660     apply (rule exE [OF continuous_on_avoid [OF contg _ \<open>g \<xi> \<noteq> 0\<close>]])
       
   661     using \<open>0 < r\<close>
       
   662     apply force
       
   663     by (metis \<open>0 < r\<close> less_trans mem_ball not_less_iff_gr_or_eq)
       
   664   show ?thesis
       
   665     apply (rule that [where g=g and r ="min r d"])
       
   666     using \<open>0 < r\<close> \<open>0 < d\<close> holg
       
   667     apply (auto simp: feq holomorphic_on_subset subset_ball d)
       
   668     done
       
   669 qed
       
   670 
       
   671 
       
   672 lemma holomorphic_factor_order_of_zero_strong:
       
   673   assumes holf: "f holomorphic_on S" "open S"  "\<xi> \<in> S" "0 < n"
       
   674       and "(deriv ^^ n) f \<xi> \<noteq> 0"
       
   675       and "\<And>i. \<lbrakk>0 < i; i < n\<rbrakk> \<Longrightarrow> (deriv ^^ i) f \<xi> = 0"
       
   676    obtains g r where "0 < r"
       
   677                 "g holomorphic_on ball \<xi> r"
       
   678                 "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = ((w - \<xi>) * g w) ^ n"
       
   679                 "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
       
   680 proof -
       
   681   obtain g r where "0 < r"
       
   682                and holg: "g holomorphic_on ball \<xi> r"
       
   683                and feq: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>)^n * g w"
       
   684                and gne: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
       
   685     by (auto intro: holomorphic_factor_order_of_zero [OF assms])
       
   686   have con: "continuous_on (ball \<xi> r) (\<lambda>z. deriv g z / g z)"
       
   687     by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on)
       
   688   have cd: "\<And>x. dist \<xi> x < r \<Longrightarrow> (\<lambda>z. deriv g z / g z) field_differentiable at x"
       
   689     apply (rule derivative_intros)+
       
   690     using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at)
       
   691     apply (metis open_ball at_within_open holg holomorphic_on_def mem_ball)
       
   692     using gne mem_ball by blast
       
   693   obtain h where h: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> (h has_field_derivative deriv g x / g x) (at x)"
       
   694     apply (rule exE [OF holomorphic_convex_primitive [of "ball \<xi> r" "{}" "\<lambda>z. deriv g z / g z"]])
       
   695     apply (auto simp: con cd)
       
   696     apply (metis open_ball at_within_open mem_ball)
       
   697     done
       
   698   then have "continuous_on (ball \<xi> r) h"
       
   699     by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open)
       
   700   then have con: "continuous_on (ball \<xi> r) (\<lambda>x. exp (h x) / g x)"
       
   701     by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne)
       
   702   have 0: "dist \<xi> x < r \<Longrightarrow> ((\<lambda>x. exp (h x) / g x) has_field_derivative 0) (at x)" for x
       
   703     apply (rule h derivative_eq_intros | simp)+
       
   704     apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2])
       
   705     using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h)
       
   706     done
       
   707   obtain c where c: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> exp (h x) / g x = c"
       
   708     by (rule DERIV_zero_connected_constant [of "ball \<xi> r" "{}" "\<lambda>x. exp(h x) / g x"]) (auto simp: con 0)
       
   709   have hol: "(\<lambda>z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \<xi> r"
       
   710     apply (rule holomorphic_on_compose [unfolded o_def, where g = exp])
       
   711     apply (rule holomorphic_intros)+
       
   712     using h holomorphic_on_open apply blast
       
   713     apply (rule holomorphic_intros)+
       
   714     using \<open>0 < n\<close> apply simp
       
   715     apply (rule holomorphic_intros)+
       
   716     done
       
   717   show ?thesis
       
   718     apply (rule that [where g="\<lambda>z. exp((Ln(inverse c) + h z)/n)" and r =r])
       
   719     using \<open>0 < r\<close> \<open>0 < n\<close>
       
   720     apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric])
       
   721     apply (rule hol)
       
   722     apply (simp add: Transcendental.exp_add gne)
       
   723     done
       
   724 qed
       
   725 
       
   726 
       
   727 lemma
       
   728   fixes k :: "'a::wellorder"
       
   729   assumes a_def: "a == LEAST x. P x" and P: "P k"
       
   730   shows def_LeastI: "P a" and def_Least_le: "a \<le> k"
       
   731 unfolding a_def
       
   732 by (rule LeastI Least_le; rule P)+
       
   733 
       
   734 lemma holomorphic_factor_zero_nonconstant:
       
   735   assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
       
   736       and "\<xi> \<in> S" "f \<xi> = 0"
       
   737       and nonconst: "\<not> f constant_on S"
       
   738    obtains g r n
       
   739       where "0 < n"  "0 < r"  "ball \<xi> r \<subseteq> S"
       
   740             "g holomorphic_on ball \<xi> r"
       
   741             "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w = (w - \<xi>)^n * g w"
       
   742             "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
       
   743 proof (cases "\<forall>n>0. (deriv ^^ n) f \<xi> = 0")
       
   744   case True then show ?thesis
       
   745     using holomorphic_fun_eq_const_on_connected [OF holf S _ \<open>\<xi> \<in> S\<close>] nonconst by (simp add: constant_on_def)
       
   746 next
       
   747   case False
       
   748   then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0" by blast
       
   749   obtain r0 where "r0 > 0" "ball \<xi> r0 \<subseteq> S" using S openE \<open>\<xi> \<in> S\<close> by auto
       
   750   define n where "n \<equiv> LEAST n. (deriv ^^ n) f \<xi> \<noteq> 0"
       
   751   have n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0"
       
   752     by (rule def_LeastI [OF n_def]) (rule n0)
       
   753   then have "0 < n" using \<open>f \<xi> = 0\<close>
       
   754     using funpow_0 by fastforce
       
   755   have n_min: "\<And>k. k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0"
       
   756     using def_Least_le [OF n_def] not_le by blast
       
   757   then obtain g r1
       
   758     where  "0 < r1" "g holomorphic_on ball \<xi> r1"
       
   759            "\<And>w. w \<in> ball \<xi> r1 \<Longrightarrow> f w = (w - \<xi>) ^ n * g w"
       
   760            "\<And>w. w \<in> ball \<xi> r1 \<Longrightarrow> g w \<noteq> 0"
       
   761     by (auto intro: holomorphic_factor_order_of_zero [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> \<open>n > 0\<close> n_ne] simp: \<open>f \<xi> = 0\<close>)
       
   762   then show ?thesis
       
   763     apply (rule_tac g=g and r="min r0 r1" and n=n in that)
       
   764     using \<open>0 < n\<close> \<open>0 < r0\<close> \<open>0 < r1\<close> \<open>ball \<xi> r0 \<subseteq> S\<close>
       
   765     apply (auto simp: subset_ball intro: holomorphic_on_subset)
       
   766     done
       
   767 qed
       
   768 
       
   769 
       
   770 lemma holomorphic_lower_bound_difference:
       
   771   assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
       
   772       and "\<xi> \<in> S" and "\<phi> \<in> S"
       
   773       and fne: "f \<phi> \<noteq> f \<xi>"
       
   774    obtains k n r
       
   775       where "0 < k"  "0 < r"
       
   776             "ball \<xi> r \<subseteq> S"
       
   777             "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> k * norm(w - \<xi>)^n \<le> norm(f w - f \<xi>)"
       
   778 proof -
       
   779   define n where "n = (LEAST n. 0 < n \<and> (deriv ^^ n) f \<xi> \<noteq> 0)"
       
   780   obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0"
       
   781     using fne holomorphic_fun_eq_const_on_connected [OF holf S] \<open>\<xi> \<in> S\<close> \<open>\<phi> \<in> S\<close> by blast
       
   782   then have "0 < n" and n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0"
       
   783     unfolding n_def by (metis (mono_tags, lifting) LeastI)+
       
   784   have n_min: "\<And>k. \<lbrakk>0 < k; k < n\<rbrakk> \<Longrightarrow> (deriv ^^ k) f \<xi> = 0"
       
   785     unfolding n_def by (blast dest: not_less_Least)
       
   786   then obtain g r
       
   787     where "0 < r" and holg: "g holomorphic_on ball \<xi> r"
       
   788       and fne: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>) ^ n * g w"
       
   789       and gnz: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
       
   790       by (auto intro: holomorphic_factor_order_of_zero  [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> \<open>n > 0\<close> n_ne])
       
   791   obtain e where "e>0" and e: "ball \<xi> e \<subseteq> S" using assms by (blast elim!: openE)
       
   792   then have holfb: "f holomorphic_on ball \<xi> e"
       
   793     using holf holomorphic_on_subset by blast
       
   794   define d where "d = (min e r) / 2"
       
   795   have "0 < d" using \<open>0 < r\<close> \<open>0 < e\<close> by (simp add: d_def)
       
   796   have "d < r"
       
   797     using \<open>0 < r\<close> by (auto simp: d_def)
       
   798   then have cbb: "cball \<xi> d \<subseteq> ball \<xi> r"
       
   799     by (auto simp: cball_subset_ball_iff)
       
   800   then have "g holomorphic_on cball \<xi> d"
       
   801     by (rule holomorphic_on_subset [OF holg])
       
   802   then have "closed (g ` cball \<xi> d)"
       
   803     by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on)
       
   804   moreover have "g ` cball \<xi> d \<noteq> {}"
       
   805     using \<open>0 < d\<close> by auto
       
   806   ultimately obtain x where x: "x \<in> g ` cball \<xi> d" and "\<And>y. y \<in> g ` cball \<xi> d \<Longrightarrow> dist 0 x \<le> dist 0 y"
       
   807     by (rule distance_attains_inf) blast
       
   808   then have leg: "\<And>w. w \<in> cball \<xi> d \<Longrightarrow> norm x \<le> norm (g w)"
       
   809     by auto
       
   810   have "ball \<xi> d \<subseteq> cball \<xi> d" by auto
       
   811   also have "... \<subseteq> ball \<xi> e" using \<open>0 < d\<close> d_def by auto
       
   812   also have "... \<subseteq> S" by (rule e)
       
   813   finally have dS: "ball \<xi> d \<subseteq> S" .
       
   814   moreover have "x \<noteq> 0" using gnz x \<open>d < r\<close> by auto
       
   815   ultimately show ?thesis
       
   816     apply (rule_tac k="norm x" and n=n and r=d in that)
       
   817     using \<open>d < r\<close> leg
       
   818     apply (auto simp: \<open>0 < d\<close> fne norm_mult norm_power algebra_simps mult_right_mono)
       
   819     done
       
   820 qed
       
   821 
       
   822 lemma
       
   823   assumes holf: "f holomorphic_on (S - {\<xi>})" and \<xi>: "\<xi> \<in> interior S"
       
   824     shows holomorphic_on_extend_lim:
       
   825           "(\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S - {\<xi>}. g z = f z)) \<longleftrightarrow>
       
   826            ((\<lambda>z. (z - \<xi>) * f z) \<longlongrightarrow> 0) (at \<xi>)"
       
   827           (is "?P = ?Q")
       
   828      and holomorphic_on_extend_bounded:
       
   829           "(\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S - {\<xi>}. g z = f z)) \<longleftrightarrow>
       
   830            (\<exists>B. eventually (\<lambda>z. norm(f z) \<le> B) (at \<xi>))"
       
   831           (is "?P = ?R")
       
   832 proof -
       
   833   obtain \<delta> where "0 < \<delta>" and \<delta>: "ball \<xi> \<delta> \<subseteq> S"
       
   834     using \<xi> mem_interior by blast
       
   835   have "?R" if holg: "g holomorphic_on S" and gf: "\<And>z. z \<in> S - {\<xi>} \<Longrightarrow> g z = f z" for g
       
   836   proof -
       
   837     have *: "\<forall>\<^sub>F z in at \<xi>. dist (g z) (g \<xi>) < 1 \<longrightarrow> cmod (f z) \<le> cmod (g \<xi>) + 1"
       
   838       apply (simp add: eventually_at)
       
   839       apply (rule_tac x="\<delta>" in exI)
       
   840       using \<delta> \<open>0 < \<delta>\<close>
       
   841       apply (clarsimp simp:)
       
   842       apply (drule_tac c=x in subsetD)
       
   843       apply (simp add: dist_commute)
       
   844       by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD)
       
   845     have "continuous_on (interior S) g"
       
   846       by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset)
       
   847     then have "\<And>x. x \<in> interior S \<Longrightarrow> (g \<longlongrightarrow> g x) (at x)"
       
   848       using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast
       
   849     then have "(g \<longlongrightarrow> g \<xi>) (at \<xi>)"
       
   850       by (simp add: \<xi>)
       
   851     then show ?thesis
       
   852       apply (rule_tac x="norm(g \<xi>) + 1" in exI)
       
   853       apply (rule eventually_mp [OF * tendstoD [where e=1]], auto)
       
   854       done
       
   855   qed
       
   856   moreover have "?Q" if "\<forall>\<^sub>F z in at \<xi>. cmod (f z) \<le> B" for B
       
   857     by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero)
       
   858   moreover have "?P" if "(\<lambda>z. (z - \<xi>) * f z) \<midarrow>\<xi>\<rightarrow> 0"
       
   859   proof -
       
   860     define h where [abs_def]: "h z = (z - \<xi>)^2 * f z" for z
       
   861     have h0: "(h has_field_derivative 0) (at \<xi>)"
       
   862       apply (simp add: h_def has_field_derivative_iff)
       
   863       apply (rule Lim_transform_within [OF that, of 1])
       
   864       apply (auto simp: field_split_simps power2_eq_square)
       
   865       done
       
   866     have holh: "h holomorphic_on S"
       
   867     proof (simp add: holomorphic_on_def, clarify)
       
   868       fix z assume "z \<in> S"
       
   869       show "h field_differentiable at z within S"
       
   870       proof (cases "z = \<xi>")
       
   871         case True then show ?thesis
       
   872           using field_differentiable_at_within field_differentiable_def h0 by blast
       
   873       next
       
   874         case False
       
   875         then have "f field_differentiable at z within S"
       
   876           using holomorphic_onD [OF holf, of z] \<open>z \<in> S\<close>
       
   877           unfolding field_differentiable_def has_field_derivative_iff
       
   878           by (force intro: exI [where x="dist \<xi> z"] elim: Lim_transform_within_set [unfolded eventually_at])
       
   879         then show ?thesis
       
   880           by (simp add: h_def power2_eq_square derivative_intros)
       
   881       qed
       
   882     qed
       
   883     define g where [abs_def]: "g z = (if z = \<xi> then deriv h \<xi> else (h z - h \<xi>) / (z - \<xi>))" for z
       
   884     have holg: "g holomorphic_on S"
       
   885       unfolding g_def by (rule pole_lemma [OF holh \<xi>])
       
   886     show ?thesis
       
   887       apply (rule_tac x="\<lambda>z. if z = \<xi> then deriv g \<xi> else (g z - g \<xi>)/(z - \<xi>)" in exI)
       
   888       apply (rule conjI)
       
   889       apply (rule pole_lemma [OF holg \<xi>])
       
   890       apply (auto simp: g_def power2_eq_square divide_simps)
       
   891       using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square)
       
   892       done
       
   893   qed
       
   894   ultimately show "?P = ?Q" and "?P = ?R"
       
   895     by meson+
       
   896 qed
       
   897 
       
   898 lemma pole_at_infinity:
       
   899   assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \<longlongrightarrow> l) at_infinity"
       
   900   obtains a n where "\<And>z. f z = (\<Sum>i\<le>n. a i * z^i)"
       
   901 proof (cases "l = 0")
       
   902   case False
       
   903   with tendsto_inverse [OF lim] show ?thesis
       
   904     apply (rule_tac a="(\<lambda>n. inverse l)" and n=0 in that)
       
   905     apply (simp add: Liouville_weak [OF holf, of "inverse l"])
       
   906     done
       
   907 next
       
   908   case True
       
   909   then have [simp]: "l = 0" .
       
   910   show ?thesis
       
   911   proof (cases "\<exists>r. 0 < r \<and> (\<forall>z \<in> ball 0 r - {0}. f(inverse z) \<noteq> 0)")
       
   912     case True
       
   913       then obtain r where "0 < r" and r: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> f(inverse z) \<noteq> 0"
       
   914              by auto
       
   915       have 1: "inverse \<circ> f \<circ> inverse holomorphic_on ball 0 r - {0}"
       
   916         by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+
       
   917       have 2: "0 \<in> interior (ball 0 r)"
       
   918         using \<open>0 < r\<close> by simp
       
   919       have "\<exists>B. 0<B \<and> eventually (\<lambda>z. cmod ((inverse \<circ> f \<circ> inverse) z) \<le> B) (at 0)"
       
   920         apply (rule exI [where x=1])
       
   921         apply simp
       
   922         using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one]
       
   923         apply (rule eventually_mono)
       
   924         apply (simp add: dist_norm)
       
   925         done
       
   926       with holomorphic_on_extend_bounded [OF 1 2]
       
   927       obtain g where holg: "g holomorphic_on ball 0 r"
       
   928                  and geq: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> g z = (inverse \<circ> f \<circ> inverse) z"
       
   929         by meson
       
   930       have ifi0: "(inverse \<circ> f \<circ> inverse) \<midarrow>0\<rightarrow> 0"
       
   931         using \<open>l = 0\<close> lim lim_at_infinity_0 by blast
       
   932       have g2g0: "g \<midarrow>0\<rightarrow> g 0"
       
   933         using \<open>0 < r\<close> centre_in_ball continuous_at continuous_on_eq_continuous_at holg
       
   934         by (blast intro: holomorphic_on_imp_continuous_on)
       
   935       have g2g1: "g \<midarrow>0\<rightarrow> 0"
       
   936         apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]])
       
   937         using \<open>0 < r\<close> by (auto simp: geq)
       
   938       have [simp]: "g 0 = 0"
       
   939         by (rule tendsto_unique [OF _ g2g0 g2g1]) simp
       
   940       have "ball 0 r - {0::complex} \<noteq> {}"
       
   941         using \<open>0 < r\<close>
       
   942         apply (clarsimp simp: ball_def dist_norm)
       
   943         apply (drule_tac c="of_real r/2" in subsetD, auto)
       
   944         done
       
   945       then obtain w::complex where "w \<noteq> 0" and w: "norm w < r" by force
       
   946       then have "g w \<noteq> 0" by (simp add: geq r)
       
   947       obtain B n e where "0 < B" "0 < e" "e \<le> r"
       
   948                      and leg: "\<And>w. norm w < e \<Longrightarrow> B * cmod w ^ n \<le> cmod (g w)"
       
   949         apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w])
       
   950         using \<open>0 < r\<close> w \<open>g w \<noteq> 0\<close> by (auto simp: ball_subset_ball_iff)
       
   951       have "cmod (f z) \<le> cmod z ^ n / B" if "2/e \<le> cmod z" for z
       
   952       proof -
       
   953         have ize: "inverse z \<in> ball 0 e - {0}" using that \<open>0 < e\<close>
       
   954           by (auto simp: norm_divide field_split_simps algebra_simps)
       
   955         then have [simp]: "z \<noteq> 0" and izr: "inverse z \<in> ball 0 r - {0}" using  \<open>e \<le> r\<close>
       
   956           by auto
       
   957         then have [simp]: "f z \<noteq> 0"
       
   958           using r [of "inverse z"] by simp
       
   959         have [simp]: "f z = inverse (g (inverse z))"
       
   960           using izr geq [of "inverse z"] by simp
       
   961         show ?thesis using ize leg [of "inverse z"]  \<open>0 < B\<close>  \<open>0 < e\<close>
       
   962           by (simp add: field_split_simps norm_divide algebra_simps)
       
   963       qed
       
   964       then show ?thesis
       
   965         apply (rule_tac a = "\<lambda>k. (deriv ^^ k) f 0 / (fact k)" and n=n in that)
       
   966         apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp)
       
   967         done
       
   968   next
       
   969     case False
       
   970     then have fi0: "\<And>r. r > 0 \<Longrightarrow> \<exists>z\<in>ball 0 r - {0}. f (inverse z) = 0"
       
   971       by simp
       
   972     have fz0: "f z = 0" if "0 < r" and lt1: "\<And>x. x \<noteq> 0 \<Longrightarrow> cmod x < r \<Longrightarrow> inverse (cmod (f (inverse x))) < 1"
       
   973               for z r
       
   974     proof -
       
   975       have f0: "(f \<longlongrightarrow> 0) at_infinity"
       
   976       proof -
       
   977         have DIM_complex[intro]: "2 \<le> DIM(complex)"  \<comment> \<open>should not be necessary!\<close>
       
   978           by simp
       
   979         from lt1 have "f (inverse x) \<noteq> 0 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> cmod x < r \<Longrightarrow> 1 < cmod (f (inverse x))" for x
       
   980           using one_less_inverse by force
       
   981         then have **: "cmod (f (inverse x)) \<le> 1 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> cmod x < r \<Longrightarrow> f (inverse x) = 0" for x
       
   982           by force
       
   983         then have *: "(f \<circ> inverse) ` (ball 0 r - {0}) \<subseteq> {0} \<union> - ball 0 1"
       
   984           by force
       
   985         have "continuous_on (inverse ` (ball 0 r - {0})) f"
       
   986           using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast
       
   987         then have "connected ((f \<circ> inverse) ` (ball 0 r - {0}))"
       
   988           apply (intro connected_continuous_image continuous_intros)
       
   989           apply (force intro: connected_punctured_ball)+
       
   990           done
       
   991         then have "{0} \<inter> (f \<circ> inverse) ` (ball 0 r - {0}) = {} \<or> - ball 0 1 \<inter> (f \<circ> inverse) ` (ball 0 r - {0}) = {}"
       
   992           by (rule connected_closedD) (use * in auto)
       
   993         then have "w \<noteq> 0 \<Longrightarrow> cmod w < r \<Longrightarrow> f (inverse w) = 0" for w
       
   994           using fi0 **[of w] \<open>0 < r\<close>
       
   995           apply (auto simp add: inf.commute [of "- ball 0 1"] Diff_eq [symmetric] image_subset_iff dest: less_imp_le)
       
   996            apply fastforce
       
   997           apply (drule bspec [of _ _ w])
       
   998            apply (auto dest: less_imp_le)
       
   999           done
       
  1000         then show ?thesis
       
  1001           apply (simp add: lim_at_infinity_0)
       
  1002           apply (rule tendsto_eventually)
       
  1003           apply (simp add: eventually_at)
       
  1004           apply (rule_tac x=r in exI)
       
  1005           apply (simp add: \<open>0 < r\<close> dist_norm)
       
  1006           done
       
  1007       qed
       
  1008       obtain w where "w \<in> ball 0 r - {0}" and "f (inverse w) = 0"
       
  1009         using False \<open>0 < r\<close> by blast
       
  1010       then show ?thesis
       
  1011         by (auto simp: f0 Liouville_weak [OF holf, of 0])
       
  1012     qed
       
  1013     show ?thesis
       
  1014       apply (rule that [of "\<lambda>n. 0" 0])
       
  1015       using lim [unfolded lim_at_infinity_0]
       
  1016       apply (simp add: Lim_at dist_norm norm_inverse)
       
  1017       apply (drule_tac x=1 in spec)
       
  1018       using fz0 apply auto
       
  1019       done
       
  1020     qed
       
  1021 qed
       
  1022 
       
  1023 subsection\<^marker>\<open>tag unimportant\<close> \<open>Entire proper functions are precisely the non-trivial polynomials\<close>
       
  1024 
       
  1025 lemma proper_map_polyfun:
       
  1026     fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}"
       
  1027   assumes "closed S" and "compact K" and c: "c i \<noteq> 0" "1 \<le> i" "i \<le> n"
       
  1028     shows "compact (S \<inter> {z. (\<Sum>i\<le>n. c i * z^i) \<in> K})"
       
  1029 proof -
       
  1030   obtain B where "B > 0" and B: "\<And>x. x \<in> K \<Longrightarrow> norm x \<le> B"
       
  1031     by (metis compact_imp_bounded \<open>compact K\<close> bounded_pos)
       
  1032   have *: "norm x \<le> b"
       
  1033             if "\<And>x. b \<le> norm x \<Longrightarrow> B + 1 \<le> norm (\<Sum>i\<le>n. c i * x ^ i)"
       
  1034                "(\<Sum>i\<le>n. c i * x ^ i) \<in> K"  for b x
       
  1035   proof -
       
  1036     have "norm (\<Sum>i\<le>n. c i * x ^ i) \<le> B"
       
  1037       using B that by blast
       
  1038     moreover have "\<not> B + 1 \<le> B"
       
  1039       by simp
       
  1040     ultimately show "norm x \<le> b"
       
  1041       using that by (metis (no_types) less_eq_real_def not_less order_trans)
       
  1042   qed
       
  1043   have "bounded {z. (\<Sum>i\<le>n. c i * z ^ i) \<in> K}"
       
  1044     using polyfun_extremal [where c=c and B="B+1", OF c]
       
  1045     by (auto simp: bounded_pos eventually_at_infinity_pos *)
       
  1046   moreover have "closed ((\<lambda>z. (\<Sum>i\<le>n. c i * z ^ i)) -` K)"
       
  1047     apply (intro allI continuous_closed_vimage continuous_intros)
       
  1048     using \<open>compact K\<close> compact_eq_bounded_closed by blast
       
  1049   ultimately show ?thesis
       
  1050     using closed_Int_compact [OF \<open>closed S\<close>] compact_eq_bounded_closed
       
  1051     by (auto simp add: vimage_def)
       
  1052 qed
       
  1053 
       
  1054 lemma proper_map_polyfun_univ:
       
  1055     fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}"
       
  1056   assumes "compact K" "c i \<noteq> 0" "1 \<le> i" "i \<le> n"
       
  1057     shows "compact ({z. (\<Sum>i\<le>n. c i * z^i) \<in> K})"
       
  1058   using proper_map_polyfun [of UNIV K c i n] assms by simp
       
  1059 
       
  1060 lemma proper_map_polyfun_eq:
       
  1061   assumes "f holomorphic_on UNIV"
       
  1062     shows "(\<forall>k. compact k \<longrightarrow> compact {z. f z \<in> k}) \<longleftrightarrow>
       
  1063            (\<exists>c n. 0 < n \<and> (c n \<noteq> 0) \<and> f = (\<lambda>z. \<Sum>i\<le>n. c i * z^i))"
       
  1064           (is "?lhs = ?rhs")
       
  1065 proof
       
  1066   assume compf [rule_format]: ?lhs
       
  1067   have 2: "\<exists>k. 0 < k \<and> a k \<noteq> 0 \<and> f = (\<lambda>z. \<Sum>i \<le> k. a i * z ^ i)"
       
  1068         if "\<And>z. f z = (\<Sum>i\<le>n. a i * z ^ i)" for a n
       
  1069   proof (cases "\<forall>i\<le>n. 0<i \<longrightarrow> a i = 0")
       
  1070     case True
       
  1071     then have [simp]: "\<And>z. f z = a 0"
       
  1072       by (simp add: that sum.atMost_shift)
       
  1073     have False using compf [of "{a 0}"] by simp
       
  1074     then show ?thesis ..
       
  1075   next
       
  1076     case False
       
  1077     then obtain k where k: "0 < k" "k\<le>n" "a k \<noteq> 0" by force
       
  1078     define m where "m = (GREATEST k. k\<le>n \<and> a k \<noteq> 0)"
       
  1079     have m: "m\<le>n \<and> a m \<noteq> 0"
       
  1080       unfolding m_def
       
  1081       apply (rule GreatestI_nat [where b = n])
       
  1082       using k apply auto
       
  1083       done
       
  1084     have [simp]: "a i = 0" if "m < i" "i \<le> n" for i
       
  1085       using Greatest_le_nat [where b = "n" and P = "\<lambda>k. k\<le>n \<and> a k \<noteq> 0"]
       
  1086       using m_def not_le that by auto
       
  1087     have "k \<le> m"
       
  1088       unfolding m_def
       
  1089       apply (rule Greatest_le_nat [where b = "n"])
       
  1090       using k apply auto
       
  1091       done
       
  1092     with k m show ?thesis
       
  1093       by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right)
       
  1094   qed
       
  1095   have "((inverse \<circ> f) \<longlongrightarrow> 0) at_infinity"
       
  1096   proof (rule Lim_at_infinityI)
       
  1097     fix e::real assume "0 < e"
       
  1098     with compf [of "cball 0 (inverse e)"]
       
  1099     show "\<exists>B. \<forall>x. B \<le> cmod x \<longrightarrow> dist ((inverse \<circ> f) x) 0 \<le> e"
       
  1100       apply simp
       
  1101       apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse)
       
  1102       apply (rule_tac x="b+1" in exI)
       
  1103       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)
       
  1104       done
       
  1105   qed
       
  1106   then show ?rhs
       
  1107     apply (rule pole_at_infinity [OF assms])
       
  1108     using 2 apply blast
       
  1109     done
       
  1110 next
       
  1111   assume ?rhs
       
  1112   then obtain c n where "0 < n" "c n \<noteq> 0" "f = (\<lambda>z. \<Sum>i\<le>n. c i * z ^ i)" by blast
       
  1113   then have "compact {z. f z \<in> k}" if "compact k" for k
       
  1114     by (auto intro: proper_map_polyfun_univ [OF that])
       
  1115   then show ?lhs by blast
       
  1116 qed
       
  1117 
       
  1118 subsection \<open>Relating invertibility and nonvanishing of derivative\<close>
       
  1119 
       
  1120 lemma has_complex_derivative_locally_injective:
       
  1121   assumes holf: "f holomorphic_on S"
       
  1122       and S: "\<xi> \<in> S" "open S"
       
  1123       and dnz: "deriv f \<xi> \<noteq> 0"
       
  1124   obtains r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)"
       
  1125 proof -
       
  1126   have *: "\<exists>d>0. \<forall>x. dist \<xi> x < d \<longrightarrow> onorm (\<lambda>v. deriv f x * v - deriv f \<xi> * v) < e" if "e > 0" for e
       
  1127   proof -
       
  1128     have contdf: "continuous_on S (deriv f)"
       
  1129       by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \<open>open S\<close>)
       
  1130     obtain \<delta> where "\<delta>>0" and \<delta>: "\<And>x. \<lbrakk>x \<in> S; dist x \<xi> \<le> \<delta>\<rbrakk> \<Longrightarrow> cmod (deriv f x - deriv f \<xi>) \<le> e/2"
       
  1131       using continuous_onE [OF contdf \<open>\<xi> \<in> S\<close>, of "e/2"] \<open>0 < e\<close>
       
  1132       by (metis dist_complex_def half_gt_zero less_imp_le)
       
  1133     obtain \<epsilon> where "\<epsilon>>0" "ball \<xi> \<epsilon> \<subseteq> S"
       
  1134       by (metis openE [OF \<open>open S\<close> \<open>\<xi> \<in> S\<close>])
       
  1135     with \<open>\<delta>>0\<close> have "\<exists>\<delta>>0. \<forall>x. dist \<xi> x < \<delta> \<longrightarrow> onorm (\<lambda>v. deriv f x * v - deriv f \<xi> * v) \<le> e/2"
       
  1136       apply (rule_tac x="min \<delta> \<epsilon>" in exI)
       
  1137       apply (intro conjI allI impI Operator_Norm.onorm_le)
       
  1138       apply simp
       
  1139       apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult)
       
  1140       apply (rule mult_right_mono [OF \<delta>])
       
  1141       apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \<delta>)
       
  1142       done
       
  1143     with \<open>e>0\<close> show ?thesis by force
       
  1144   qed
       
  1145   have "inj ((*) (deriv f \<xi>))"
       
  1146     using dnz by simp
       
  1147   then obtain g' where g': "linear g'" "g' \<circ> (*) (deriv f \<xi>) = id"
       
  1148     using linear_injective_left_inverse [of "(*) (deriv f \<xi>)"]
       
  1149     by (auto simp: linear_times)
       
  1150   show ?thesis
       
  1151     apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\<lambda>z h. deriv f z * h" and g' = g'])
       
  1152     using g' *
       
  1153     apply (simp_all add: linear_conv_bounded_linear that)
       
  1154     using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf
       
  1155         holomorphic_on_imp_differentiable_at \<open>open S\<close> apply blast
       
  1156     done
       
  1157 qed
       
  1158 
       
  1159 lemma has_complex_derivative_locally_invertible:
       
  1160   assumes holf: "f holomorphic_on S"
       
  1161       and S: "\<xi> \<in> S" "open S"
       
  1162       and dnz: "deriv f \<xi> \<noteq> 0"
       
  1163   obtains r where "r > 0" "ball \<xi> r \<subseteq> S" "open (f `  (ball \<xi> r))" "inj_on f (ball \<xi> r)"
       
  1164 proof -
       
  1165   obtain r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)"
       
  1166     by (blast intro: that has_complex_derivative_locally_injective [OF assms])
       
  1167   then have \<xi>: "\<xi> \<in> ball \<xi> r" by simp
       
  1168   then have nc: "\<not> f constant_on ball \<xi> r"
       
  1169     using \<open>inj_on f (ball \<xi> r)\<close> injective_not_constant by fastforce
       
  1170   have holf': "f holomorphic_on ball \<xi> r"
       
  1171     using \<open>ball \<xi> r \<subseteq> S\<close> holf holomorphic_on_subset by blast
       
  1172   have "open (f ` ball \<xi> r)"
       
  1173     apply (rule open_mapping_thm [OF holf'])
       
  1174     using nc apply auto
       
  1175     done
       
  1176   then show ?thesis
       
  1177     using \<open>0 < r\<close> \<open>ball \<xi> r \<subseteq> S\<close> \<open>inj_on f (ball \<xi> r)\<close> that  by blast
       
  1178 qed
       
  1179 
       
  1180 lemma holomorphic_injective_imp_regular:
       
  1181   assumes holf: "f holomorphic_on S"
       
  1182       and "open S" and injf: "inj_on f S"
       
  1183       and "\<xi> \<in> S"
       
  1184     shows "deriv f \<xi> \<noteq> 0"
       
  1185 proof -
       
  1186   obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE)
       
  1187   have holf': "f holomorphic_on ball \<xi> r"
       
  1188     using \<open>ball \<xi> r \<subseteq> S\<close> holf holomorphic_on_subset by blast
       
  1189   show ?thesis
       
  1190   proof (cases "\<forall>n>0. (deriv ^^ n) f \<xi> = 0")
       
  1191     case True
       
  1192     have fcon: "f w = f \<xi>" if "w \<in> ball \<xi> r" for w
       
  1193       apply (rule holomorphic_fun_eq_const_on_connected [OF holf'])
       
  1194       using True \<open>0 < r\<close> that by auto
       
  1195     have False
       
  1196       using fcon [of "\<xi> + r/2"] \<open>0 < r\<close> r injf unfolding inj_on_def
       
  1197       by (metis \<open>\<xi> \<in> S\<close> contra_subsetD dist_commute fcon mem_ball perfect_choose_dist)
       
  1198     then show ?thesis ..
       
  1199   next
       
  1200     case False
       
  1201     then obtain n0 where n0: "n0 > 0 \<and> (deriv ^^ n0) f \<xi> \<noteq> 0" by blast
       
  1202     define n where [abs_def]: "n = (LEAST n. n > 0 \<and> (deriv ^^ n) f \<xi> \<noteq> 0)"
       
  1203     have n_ne: "n > 0" "(deriv ^^ n) f \<xi> \<noteq> 0"
       
  1204       using def_LeastI [OF n_def n0] by auto
       
  1205     have n_min: "\<And>k. 0 < k \<Longrightarrow> k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0"
       
  1206       using def_Least_le [OF n_def] not_le by auto
       
  1207     obtain g \<delta> where "0 < \<delta>"
       
  1208              and holg: "g holomorphic_on ball \<xi> \<delta>"
       
  1209              and fd: "\<And>w. w \<in> ball \<xi> \<delta> \<Longrightarrow> f w - f \<xi> = ((w - \<xi>) * g w) ^ n"
       
  1210              and gnz: "\<And>w. w \<in> ball \<xi> \<delta> \<Longrightarrow> g w \<noteq> 0"
       
  1211       apply (rule holomorphic_factor_order_of_zero_strong [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> n_ne])
       
  1212       apply (blast intro: n_min)+
       
  1213       done
       
  1214     show ?thesis
       
  1215     proof (cases "n=1")
       
  1216       case True
       
  1217       with n_ne show ?thesis by auto
       
  1218     next
       
  1219       case False
       
  1220       have holgw: "(\<lambda>w. (w - \<xi>) * g w) holomorphic_on ball \<xi> (min r \<delta>)"
       
  1221         apply (rule holomorphic_intros)+
       
  1222         using holg by (simp add: holomorphic_on_subset subset_ball)
       
  1223       have gd: "\<And>w. dist \<xi> w < \<delta> \<Longrightarrow> (g has_field_derivative deriv g w) (at w)"
       
  1224         using holg
       
  1225         by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH)
       
  1226       have *: "\<And>w. w \<in> ball \<xi> (min r \<delta>)
       
  1227             \<Longrightarrow> ((\<lambda>w. (w - \<xi>) * g w) has_field_derivative ((w - \<xi>) * deriv g w + g w))
       
  1228                 (at w)"
       
  1229         by (rule gd derivative_eq_intros | simp)+
       
  1230       have [simp]: "deriv (\<lambda>w. (w - \<xi>) * g w) \<xi> \<noteq> 0"
       
  1231         using * [of \<xi>] \<open>0 < \<delta>\<close> \<open>0 < r\<close> by (simp add: DERIV_imp_deriv gnz)
       
  1232       obtain T where "\<xi> \<in> T" "open T" and Tsb: "T \<subseteq> ball \<xi> (min r \<delta>)" and oimT: "open ((\<lambda>w. (w - \<xi>) * g w) ` T)"
       
  1233         apply (rule has_complex_derivative_locally_invertible [OF holgw, of \<xi>])
       
  1234         using \<open>0 < r\<close> \<open>0 < \<delta>\<close>
       
  1235         apply (simp_all add:)
       
  1236         by (meson open_ball centre_in_ball)
       
  1237       define U where "U = (\<lambda>w. (w - \<xi>) * g w) ` T"
       
  1238       have "open U" by (metis oimT U_def)
       
  1239       have "0 \<in> U"
       
  1240         apply (auto simp: U_def)
       
  1241         apply (rule image_eqI [where x = \<xi>])
       
  1242         apply (auto simp: \<open>\<xi> \<in> T\<close>)
       
  1243         done
       
  1244       then obtain \<epsilon> where "\<epsilon>>0" and \<epsilon>: "cball 0 \<epsilon> \<subseteq> U"
       
  1245         using \<open>open U\<close> open_contains_cball by blast
       
  1246       then have "\<epsilon> * exp(2 * of_real pi * \<i> * (0/n)) \<in> cball 0 \<epsilon>"
       
  1247                 "\<epsilon> * exp(2 * of_real pi * \<i> * (1/n)) \<in> cball 0 \<epsilon>"
       
  1248         by (auto simp: norm_mult)
       
  1249       with \<epsilon> have "\<epsilon> * exp(2 * of_real pi * \<i> * (0/n)) \<in> U"
       
  1250                   "\<epsilon> * exp(2 * of_real pi * \<i> * (1/n)) \<in> U" by blast+
       
  1251       then obtain y0 y1 where "y0 \<in> T" and y0: "(y0 - \<xi>) * g y0 = \<epsilon> * exp(2 * of_real pi * \<i> * (0/n))"
       
  1252                           and "y1 \<in> T" and y1: "(y1 - \<xi>) * g y1 = \<epsilon> * exp(2 * of_real pi * \<i> * (1/n))"
       
  1253         by (auto simp: U_def)
       
  1254       then have "y0 \<in> ball \<xi> \<delta>" "y1 \<in> ball \<xi> \<delta>" using Tsb by auto
       
  1255       moreover have "y0 \<noteq> y1"
       
  1256         using y0 y1 \<open>\<epsilon> > 0\<close> complex_root_unity_eq_1 [of n 1] \<open>n > 0\<close> False by auto
       
  1257       moreover have "T \<subseteq> S"
       
  1258         by (meson Tsb min.cobounded1 order_trans r subset_ball)
       
  1259       ultimately have False
       
  1260         using inj_onD [OF injf, of y0 y1] \<open>y0 \<in> T\<close> \<open>y1 \<in> T\<close>
       
  1261         using fd [of y0] fd [of y1] complex_root_unity [of n 1] n_ne
       
  1262         apply (simp add: y0 y1 power_mult_distrib)
       
  1263         apply (force simp: algebra_simps)
       
  1264         done
       
  1265       then show ?thesis ..
       
  1266     qed
       
  1267   qed
       
  1268 qed
       
  1269 
       
  1270 text\<open>Hence a nice clean inverse function theorem\<close>
       
  1271 
       
  1272 lemma has_field_derivative_inverse_strong:
       
  1273   fixes f :: "'a::{euclidean_space,real_normed_field} \<Rightarrow> 'a"
       
  1274   shows "\<lbrakk>DERIV f x :> f'; f' \<noteq> 0; open S; x \<in> S; continuous_on S f;
       
  1275          \<And>z. z \<in> S \<Longrightarrow> g (f z) = z\<rbrakk>
       
  1276          \<Longrightarrow> DERIV g (f x) :> inverse (f')"
       
  1277   unfolding has_field_derivative_def
       
  1278   by (rule has_derivative_inverse_strong [of S x f g]) auto
       
  1279 
       
  1280 lemma has_field_derivative_inverse_strong_x:
       
  1281   fixes f :: "'a::{euclidean_space,real_normed_field} \<Rightarrow> 'a"
       
  1282   shows  "\<lbrakk>DERIV f (g y) :> f'; f' \<noteq> 0; open S; continuous_on S f; g y \<in> S; f(g y) = y;
       
  1283            \<And>z. z \<in> S \<Longrightarrow> g (f z) = z\<rbrakk>
       
  1284           \<Longrightarrow> DERIV g y :> inverse (f')"
       
  1285   unfolding has_field_derivative_def
       
  1286   by (rule has_derivative_inverse_strong_x [of S g y f]) auto
       
  1287 
       
  1288 proposition holomorphic_has_inverse:
       
  1289   assumes holf: "f holomorphic_on S"
       
  1290       and "open S" and injf: "inj_on f S"
       
  1291   obtains g where "g holomorphic_on (f ` S)"
       
  1292                   "\<And>z. z \<in> S \<Longrightarrow> deriv f z * deriv g (f z) = 1"
       
  1293                   "\<And>z. z \<in> S \<Longrightarrow> g(f z) = z"
       
  1294 proof -
       
  1295   have ofs: "open (f ` S)"
       
  1296     by (rule open_mapping_thm3 [OF assms])
       
  1297   have contf: "continuous_on S f"
       
  1298     by (simp add: holf holomorphic_on_imp_continuous_on)
       
  1299   have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \<in> S" for z
       
  1300   proof -
       
  1301     have 1: "(f has_field_derivative deriv f z) (at z)"
       
  1302       using DERIV_deriv_iff_field_differentiable \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_on_imp_differentiable_at
       
  1303       by blast
       
  1304     have 2: "deriv f z \<noteq> 0"
       
  1305       using \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_injective_imp_regular injf by blast
       
  1306     show ?thesis
       
  1307       apply (rule has_field_derivative_inverse_strong [OF 1 2 \<open>open S\<close> \<open>z \<in> S\<close>])
       
  1308        apply (simp add: holf holomorphic_on_imp_continuous_on)
       
  1309       by (simp add: injf the_inv_into_f_f)
       
  1310   qed
       
  1311   show ?thesis
       
  1312     proof
       
  1313       show "the_inv_into S f holomorphic_on f ` S"
       
  1314         by (simp add: holomorphic_on_open ofs) (blast intro: *)
       
  1315     next
       
  1316       fix z assume "z \<in> S"
       
  1317       have "deriv f z \<noteq> 0"
       
  1318         using \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_injective_imp_regular injf by blast
       
  1319       then show "deriv f z * deriv (the_inv_into S f) (f z) = 1"
       
  1320         using * [OF \<open>z \<in> S\<close>]  by (simp add: DERIV_imp_deriv)
       
  1321     next
       
  1322       fix z assume "z \<in> S"
       
  1323       show "the_inv_into S f (f z) = z"
       
  1324         by (simp add: \<open>z \<in> S\<close> injf the_inv_into_f_f)
       
  1325   qed
       
  1326 qed
       
  1327 
       
  1328 subsection\<open>The Schwarz Lemma\<close>
       
  1329 
       
  1330 lemma Schwarz1:
       
  1331   assumes holf: "f holomorphic_on S"
       
  1332       and contf: "continuous_on (closure S) f"
       
  1333       and S: "open S" "connected S"
       
  1334       and boS: "bounded S"
       
  1335       and "S \<noteq> {}"
       
  1336   obtains w where "w \<in> frontier S"
       
  1337        "\<And>z. z \<in> closure S \<Longrightarrow> norm (f z) \<le> norm (f w)"
       
  1338 proof -
       
  1339   have connf: "continuous_on (closure S) (norm o f)"
       
  1340     using contf continuous_on_compose continuous_on_norm_id by blast
       
  1341   have coc: "compact (closure S)"
       
  1342     by (simp add: \<open>bounded S\<close> bounded_closure compact_eq_bounded_closed)
       
  1343   then obtain x where x: "x \<in> closure S" and xmax: "\<And>z. z \<in> closure S \<Longrightarrow> norm(f z) \<le> norm(f x)"
       
  1344     apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]])
       
  1345     using \<open>S \<noteq> {}\<close> apply auto
       
  1346     done
       
  1347   then show ?thesis
       
  1348   proof (cases "x \<in> frontier S")
       
  1349     case True
       
  1350     then show ?thesis using that xmax by blast
       
  1351   next
       
  1352     case False
       
  1353     then have "x \<in> S"
       
  1354       using \<open>open S\<close> frontier_def interior_eq x by auto
       
  1355     then have "f constant_on S"
       
  1356       apply (rule maximum_modulus_principle [OF holf S \<open>open S\<close> order_refl])
       
  1357       using closure_subset apply (blast intro: xmax)
       
  1358       done
       
  1359     then have "f constant_on (closure S)"
       
  1360       by (rule constant_on_closureI [OF _ contf])
       
  1361     then obtain c where c: "\<And>x. x \<in> closure S \<Longrightarrow> f x = c"
       
  1362       by (meson constant_on_def)
       
  1363     obtain w where "w \<in> frontier S"
       
  1364       by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV)
       
  1365     then show ?thesis
       
  1366       by (simp add: c frontier_def that)
       
  1367   qed
       
  1368 qed
       
  1369 
       
  1370 lemma Schwarz2:
       
  1371  "\<lbrakk>f holomorphic_on ball 0 r;
       
  1372     0 < s; ball w s \<subseteq> ball 0 r;
       
  1373     \<And>z. norm (w-z) < s \<Longrightarrow> norm(f z) \<le> norm(f w)\<rbrakk>
       
  1374     \<Longrightarrow> f constant_on ball 0 r"
       
  1375 by (rule maximum_modulus_principle [where U = "ball w s" and \<xi> = w]) (simp_all add: dist_norm)
       
  1376 
       
  1377 lemma Schwarz3:
       
  1378   assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0"
       
  1379   obtains h where "h holomorphic_on (ball 0 r)" and "\<And>z. norm z < r \<Longrightarrow> f z = z * (h z)" and "deriv f 0 = h 0"
       
  1380 proof -
       
  1381   define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z
       
  1382   have d0: "deriv f 0 = h 0"
       
  1383     by (simp add: h_def)
       
  1384   moreover have "h holomorphic_on (ball 0 r)"
       
  1385     by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def)
       
  1386   moreover have "norm z < r \<Longrightarrow> f z = z * h z" for z
       
  1387     by (simp add: h_def)
       
  1388   ultimately show ?thesis
       
  1389     using that by blast
       
  1390 qed
       
  1391 
       
  1392 proposition Schwarz_Lemma:
       
  1393   assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0"
       
  1394       and no: "\<And>z. norm z < 1 \<Longrightarrow> norm (f z) < 1"
       
  1395       and \<xi>: "norm \<xi> < 1"
       
  1396     shows "norm (f \<xi>) \<le> norm \<xi>" and "norm(deriv f 0) \<le> 1"
       
  1397       and "((\<exists>z. norm z < 1 \<and> z \<noteq> 0 \<and> norm(f z) = norm z)
       
  1398             \<or> norm(deriv f 0) = 1)
       
  1399            \<Longrightarrow> \<exists>\<alpha>. (\<forall>z. norm z < 1 \<longrightarrow> f z = \<alpha> * z) \<and> norm \<alpha> = 1"
       
  1400       (is "?P \<Longrightarrow> ?Q")
       
  1401 proof -
       
  1402   obtain h where holh: "h holomorphic_on (ball 0 1)"
       
  1403              and fz_eq: "\<And>z. norm z < 1 \<Longrightarrow> f z = z * (h z)" and df0: "deriv f 0 = h 0"
       
  1404     by (rule Schwarz3 [OF holf]) auto
       
  1405   have noh_le: "norm (h z) \<le> 1" if z: "norm z < 1" for z
       
  1406   proof -
       
  1407     have "norm (h z) < a" if a: "1 < a" for a
       
  1408     proof -
       
  1409       have "max (inverse a) (norm z) < 1"
       
  1410         using z a by (simp_all add: inverse_less_1_iff)
       
  1411       then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1"
       
  1412         using Rats_dense_in_real by blast
       
  1413       then have nzr: "norm z < r" and ira: "inverse r < a"
       
  1414         using z a less_imp_inverse_less by force+
       
  1415       then have "0 < r"
       
  1416         by (meson norm_not_less_zero not_le order.strict_trans2)
       
  1417       have holh': "h holomorphic_on ball 0 r"
       
  1418         by (meson holh \<open>r < 1\<close> holomorphic_on_subset less_eq_real_def subset_ball)
       
  1419       have conth': "continuous_on (cball 0 r) h"
       
  1420         by (meson \<open>r < 1\<close> dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI)
       
  1421       obtain w where w: "norm w = r" and lenw: "\<And>z. norm z < r \<Longrightarrow> norm(h z) \<le> norm(h w)"
       
  1422         apply (rule Schwarz1 [OF holh']) using conth' \<open>0 < r\<close> by auto
       
  1423       have "h w = f w / w" using fz_eq \<open>r < 1\<close> nzr w by auto
       
  1424       then have "cmod (h z) < inverse r"
       
  1425         by (metis \<open>0 < r\<close> \<open>r < 1\<close> divide_strict_right_mono inverse_eq_divide
       
  1426                   le_less_trans lenw no norm_divide nzr w)
       
  1427       then show ?thesis using ira by linarith
       
  1428     qed
       
  1429     then show "norm (h z) \<le> 1"
       
  1430       using not_le by blast
       
  1431   qed
       
  1432   show "cmod (f \<xi>) \<le> cmod \<xi>"
       
  1433   proof (cases "\<xi> = 0")
       
  1434     case True then show ?thesis by auto
       
  1435   next
       
  1436     case False
       
  1437     then show ?thesis
       
  1438       by (simp add: noh_le fz_eq \<xi> mult_left_le norm_mult)
       
  1439   qed
       
  1440   show no_df0: "norm(deriv f 0) \<le> 1"
       
  1441     by (simp add: \<open>\<And>z. cmod z < 1 \<Longrightarrow> cmod (h z) \<le> 1\<close> df0)
       
  1442   show "?Q" if "?P"
       
  1443     using that
       
  1444   proof
       
  1445     assume "\<exists>z. cmod z < 1 \<and> z \<noteq> 0 \<and> cmod (f z) = cmod z"
       
  1446     then obtain \<gamma> where \<gamma>: "cmod \<gamma> < 1" "\<gamma> \<noteq> 0" "cmod (f \<gamma>) = cmod \<gamma>" by blast
       
  1447     then have [simp]: "norm (h \<gamma>) = 1"
       
  1448       by (simp add: fz_eq norm_mult)
       
  1449     have "ball \<gamma> (1 - cmod \<gamma>) \<subseteq> ball 0 1"
       
  1450       by (simp add: ball_subset_ball_iff)
       
  1451     moreover have "\<And>z. cmod (\<gamma> - z) < 1 - cmod \<gamma> \<Longrightarrow> cmod (h z) \<le> cmod (h \<gamma>)"
       
  1452       apply (simp add: algebra_simps)
       
  1453       by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4)
       
  1454     ultimately obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c"
       
  1455       using Schwarz2 [OF holh, of "1 - norm \<gamma>" \<gamma>, unfolded constant_on_def] \<gamma> by auto
       
  1456     then have "norm c = 1"
       
  1457       using \<gamma> by force
       
  1458     with c show ?thesis
       
  1459       using fz_eq by auto
       
  1460   next
       
  1461     assume [simp]: "cmod (deriv f 0) = 1"
       
  1462     then obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c"
       
  1463       using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le
       
  1464       by auto
       
  1465     moreover have "norm c = 1"  using df0 c by auto
       
  1466     ultimately show ?thesis
       
  1467       using fz_eq by auto
       
  1468   qed
       
  1469 qed
       
  1470 
       
  1471 corollary Schwarz_Lemma':
       
  1472   assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0"
       
  1473       and no: "\<And>z. norm z < 1 \<Longrightarrow> norm (f z) < 1"
       
  1474     shows "((\<forall>\<xi>. norm \<xi> < 1 \<longrightarrow> norm (f \<xi>) \<le> norm \<xi>)
       
  1475             \<and> norm(deriv f 0) \<le> 1)
       
  1476             \<and> (((\<exists>z. norm z < 1 \<and> z \<noteq> 0 \<and> norm(f z) = norm z)
       
  1477               \<or> norm(deriv f 0) = 1)
       
  1478               \<longrightarrow> (\<exists>\<alpha>. (\<forall>z. norm z < 1 \<longrightarrow> f z = \<alpha> * z) \<and> norm \<alpha> = 1))"
       
  1479   using Schwarz_Lemma [OF assms]
       
  1480   by (metis (no_types) norm_eq_zero zero_less_one)
       
  1481 
       
  1482 subsection\<open>The Schwarz reflection principle\<close>
       
  1483 
       
  1484 lemma hol_pal_lem0:
       
  1485   assumes "d \<bullet> a \<le> k" "k \<le> d \<bullet> b"
       
  1486   obtains c where
       
  1487      "c \<in> closed_segment a b" "d \<bullet> c = k"
       
  1488      "\<And>z. z \<in> closed_segment a c \<Longrightarrow> d \<bullet> z \<le> k"
       
  1489      "\<And>z. z \<in> closed_segment c b \<Longrightarrow> k \<le> d \<bullet> z"
       
  1490 proof -
       
  1491   obtain c where cin: "c \<in> closed_segment a b" and keq: "k = d \<bullet> c"
       
  1492     using connected_ivt_hyperplane [of "closed_segment a b" a b d k]
       
  1493     by (auto simp: assms)
       
  1494   have "closed_segment a c \<subseteq> {z. d \<bullet> z \<le> k}"  "closed_segment c b \<subseteq> {z. k \<le> d \<bullet> z}"
       
  1495     unfolding segment_convex_hull using assms keq
       
  1496     by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal)
       
  1497   then show ?thesis using cin that by fastforce
       
  1498 qed
       
  1499 
       
  1500 lemma hol_pal_lem1:
       
  1501   assumes "convex S" "open S"
       
  1502       and abc: "a \<in> S" "b \<in> S" "c \<in> S"
       
  1503           "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" "d \<bullet> b \<le> k" "d \<bullet> c \<le> k"
       
  1504       and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}"
       
  1505       and contf: "continuous_on S f"
       
  1506     shows "contour_integral (linepath a b) f +
       
  1507            contour_integral (linepath b c) f +
       
  1508            contour_integral (linepath c a) f = 0"
       
  1509 proof -
       
  1510   have "interior (convex hull {a, b, c}) \<subseteq> interior(S \<inter> {x. d \<bullet> x \<le> k})"
       
  1511     apply (rule interior_mono)
       
  1512     apply (rule hull_minimal)
       
  1513      apply (simp add: abc lek)
       
  1514     apply (rule convex_Int [OF \<open>convex S\<close> convex_halfspace_le])
       
  1515     done
       
  1516   also have "... \<subseteq> {z \<in> S. d \<bullet> z < k}"
       
  1517     by (force simp: interior_open [OF \<open>open S\<close>] \<open>d \<noteq> 0\<close>)
       
  1518   finally have *: "interior (convex hull {a, b, c}) \<subseteq> {z \<in> S. d \<bullet> z < k}" .
       
  1519   have "continuous_on (convex hull {a,b,c}) f"
       
  1520     using \<open>convex S\<close> contf abc continuous_on_subset subset_hull
       
  1521     by fastforce
       
  1522   moreover have "f holomorphic_on interior (convex hull {a,b,c})"
       
  1523     by (rule holomorphic_on_subset [OF holf1 *])
       
  1524   ultimately show ?thesis
       
  1525     using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3
       
  1526       by blast
       
  1527 qed
       
  1528 
       
  1529 lemma hol_pal_lem2:
       
  1530   assumes S: "convex S" "open S"
       
  1531       and abc: "a \<in> S" "b \<in> S" "c \<in> S"
       
  1532       and "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" "d \<bullet> b \<le> k"
       
  1533       and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}"
       
  1534       and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}"
       
  1535       and contf: "continuous_on S f"
       
  1536     shows "contour_integral (linepath a b) f +
       
  1537            contour_integral (linepath b c) f +
       
  1538            contour_integral (linepath c a) f = 0"
       
  1539 proof (cases "d \<bullet> c \<le> k")
       
  1540   case True show ?thesis
       
  1541     by (rule hol_pal_lem1 [OF S abc \<open>d \<noteq> 0\<close> lek True holf1 contf])
       
  1542 next
       
  1543   case False
       
  1544   then have "d \<bullet> c > k" by force
       
  1545   obtain a' where a': "a' \<in> closed_segment b c" and "d \<bullet> a' = k"
       
  1546      and ba': "\<And>z. z \<in> closed_segment b a' \<Longrightarrow> d \<bullet> z \<le> k"
       
  1547      and a'c: "\<And>z. z \<in> closed_segment a' c \<Longrightarrow> k \<le> d \<bullet> z"
       
  1548     apply (rule hol_pal_lem0 [of d b k c, OF \<open>d \<bullet> b \<le> k\<close>])
       
  1549     using False by auto
       
  1550   obtain b' where b': "b' \<in> closed_segment a c" and "d \<bullet> b' = k"
       
  1551      and ab': "\<And>z. z \<in> closed_segment a b' \<Longrightarrow> d \<bullet> z \<le> k"
       
  1552      and b'c: "\<And>z. z \<in> closed_segment b' c \<Longrightarrow> k \<le> d \<bullet> z"
       
  1553     apply (rule hol_pal_lem0 [of d a k c, OF \<open>d \<bullet> a \<le> k\<close>])
       
  1554     using False by auto
       
  1555   have a'b': "a' \<in> S \<and> b' \<in> S"
       
  1556     using a' abc b' convex_contains_segment \<open>convex S\<close> by auto
       
  1557   have "continuous_on (closed_segment c a) f"
       
  1558     by (meson abc contf continuous_on_subset convex_contains_segment \<open>convex S\<close>)
       
  1559   then have 1: "contour_integral (linepath c a) f =
       
  1560                 contour_integral (linepath c b') f + contour_integral (linepath b' a) f"
       
  1561     apply (rule contour_integral_split_linepath)
       
  1562     using b' by (simp add: closed_segment_commute)
       
  1563   have "continuous_on (closed_segment b c) f"
       
  1564     by (meson abc contf continuous_on_subset convex_contains_segment \<open>convex S\<close>)
       
  1565   then have 2: "contour_integral (linepath b c) f =
       
  1566                 contour_integral (linepath b a') f + contour_integral (linepath a' c) f"
       
  1567     by (rule contour_integral_split_linepath [OF _ a'])
       
  1568   have 3: "contour_integral (reversepath (linepath b' a')) f =
       
  1569                 - contour_integral (linepath b' a') f"
       
  1570     by (rule contour_integral_reversepath [OF valid_path_linepath])
       
  1571   have fcd_le: "f field_differentiable at x"
       
  1572                if "x \<in> interior S \<and> x \<in> interior {x. d \<bullet> x \<le> k}" for x
       
  1573   proof -
       
  1574     have "f holomorphic_on S \<inter> {c. d \<bullet> c < k}"
       
  1575       by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1)
       
  1576     then have "\<exists>C D. x \<in> interior C \<inter> interior D \<and> f holomorphic_on interior C \<inter> interior D"
       
  1577       using that
       
  1578       by (metis Collect_mem_eq Int_Collect \<open>d \<noteq> 0\<close> interior_halfspace_le interior_open \<open>open S\<close>)
       
  1579     then show "f field_differentiable at x"
       
  1580       by (metis at_within_interior holomorphic_on_def interior_Int interior_interior)
       
  1581   qed
       
  1582   have ab_le: "\<And>x. x \<in> closed_segment a b \<Longrightarrow> d \<bullet> x \<le> k"
       
  1583   proof -
       
  1584     fix x :: complex
       
  1585     assume "x \<in> closed_segment a b"
       
  1586     then have "\<And>C. x \<in> C \<or> b \<notin> C \<or> a \<notin> C \<or> \<not> convex C"
       
  1587       by (meson contra_subsetD convex_contains_segment)
       
  1588     then show "d \<bullet> x \<le> k"
       
  1589       by (metis lek convex_halfspace_le mem_Collect_eq)
       
  1590   qed
       
  1591   have "continuous_on (S \<inter> {x. d \<bullet> x \<le> k}) f" using contf
       
  1592     by (simp add: continuous_on_subset)
       
  1593   then have "(f has_contour_integral 0)
       
  1594          (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)"
       
  1595     apply (rule Cauchy_theorem_convex [where K = "{}"])
       
  1596     apply (simp_all add: path_image_join convex_Int convex_halfspace_le \<open>convex S\<close> fcd_le ab_le
       
  1597                 closed_segment_subset abc a'b' ba')
       
  1598     by (metis \<open>d \<bullet> a' = k\<close> \<open>d \<bullet> b' = k\<close> convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl)
       
  1599   then have 4: "contour_integral (linepath a b) f +
       
  1600                 contour_integral (linepath b a') f +
       
  1601                 contour_integral (linepath a' b') f +
       
  1602                 contour_integral (linepath b' a) f = 0"
       
  1603     by (rule has_chain_integral_chain_integral4)
       
  1604   have fcd_ge: "f field_differentiable at x"
       
  1605                if "x \<in> interior S \<and> x \<in> interior {x. k \<le> d \<bullet> x}" for x
       
  1606   proof -
       
  1607     have f2: "f holomorphic_on S \<inter> {c. k < d \<bullet> c}"
       
  1608       by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2)
       
  1609     have f3: "interior S = S"
       
  1610       by (simp add: interior_open \<open>open S\<close>)
       
  1611     then have "x \<in> S \<inter> interior {c. k \<le> d \<bullet> c}"
       
  1612       using that by simp
       
  1613     then show "f field_differentiable at x"
       
  1614       using f3 f2 unfolding holomorphic_on_def
       
  1615       by (metis (no_types) \<open>d \<noteq> 0\<close> at_within_interior interior_Int interior_halfspace_ge interior_interior)
       
  1616   qed
       
  1617   have "continuous_on (S \<inter> {x. k \<le> d \<bullet> x}) f" using contf
       
  1618     by (simp add: continuous_on_subset)
       
  1619   then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')"
       
  1620     apply (rule Cauchy_theorem_convex [where K = "{}"])
       
  1621     apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \<open>convex S\<close>
       
  1622                       fcd_ge closed_segment_subset abc a'b' a'c)
       
  1623     by (metis \<open>d \<bullet> a' = k\<close> b'c closed_segment_commute convex_contains_segment
       
  1624               convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl)
       
  1625   then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0"
       
  1626     by (rule has_chain_integral_chain_integral3)
       
  1627   show ?thesis
       
  1628     using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath)
       
  1629 qed
       
  1630 
       
  1631 lemma hol_pal_lem3:
       
  1632   assumes S: "convex S" "open S"
       
  1633       and abc: "a \<in> S" "b \<in> S" "c \<in> S"
       
  1634       and "d \<noteq> 0" and lek: "d \<bullet> a \<le> k"
       
  1635       and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}"
       
  1636       and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}"
       
  1637       and contf: "continuous_on S f"
       
  1638     shows "contour_integral (linepath a b) f +
       
  1639            contour_integral (linepath b c) f +
       
  1640            contour_integral (linepath c a) f = 0"
       
  1641 proof (cases "d \<bullet> b \<le> k")
       
  1642   case True show ?thesis
       
  1643     by (rule hol_pal_lem2 [OF S abc \<open>d \<noteq> 0\<close> lek True holf1 holf2 contf])
       
  1644 next
       
  1645   case False
       
  1646   show ?thesis
       
  1647   proof (cases "d \<bullet> c \<le> k")
       
  1648     case True
       
  1649     have "contour_integral (linepath c a) f +
       
  1650           contour_integral (linepath a b) f +
       
  1651           contour_integral (linepath b c) f = 0"
       
  1652       by (rule hol_pal_lem2 [OF S \<open>c \<in> S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close> \<open>d \<noteq> 0\<close> \<open>d \<bullet> c \<le> k\<close> lek holf1 holf2 contf])
       
  1653     then show ?thesis
       
  1654       by (simp add: algebra_simps)
       
  1655   next
       
  1656     case False
       
  1657     have "contour_integral (linepath b c) f +
       
  1658           contour_integral (linepath c a) f +
       
  1659           contour_integral (linepath a b) f = 0"
       
  1660       apply (rule hol_pal_lem2 [OF S \<open>b \<in> S\<close> \<open>c \<in> S\<close> \<open>a \<in> S\<close>, of "-d" "-k"])
       
  1661       using \<open>d \<noteq> 0\<close> \<open>\<not> d \<bullet> b \<le> k\<close> False by (simp_all add: holf1 holf2 contf)
       
  1662     then show ?thesis
       
  1663       by (simp add: algebra_simps)
       
  1664   qed
       
  1665 qed
       
  1666 
       
  1667 lemma hol_pal_lem4:
       
  1668   assumes S: "convex S" "open S"
       
  1669       and abc: "a \<in> S" "b \<in> S" "c \<in> S" and "d \<noteq> 0"
       
  1670       and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}"
       
  1671       and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}"
       
  1672       and contf: "continuous_on S f"
       
  1673     shows "contour_integral (linepath a b) f +
       
  1674            contour_integral (linepath b c) f +
       
  1675            contour_integral (linepath c a) f = 0"
       
  1676 proof (cases "d \<bullet> a \<le> k")
       
  1677   case True show ?thesis
       
  1678     by (rule hol_pal_lem3 [OF S abc \<open>d \<noteq> 0\<close> True holf1 holf2 contf])
       
  1679 next
       
  1680   case False
       
  1681   show ?thesis
       
  1682     apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"])
       
  1683     using \<open>d \<noteq> 0\<close> False by (simp_all add: holf1 holf2 contf)
       
  1684 qed
       
  1685 
       
  1686 lemma holomorphic_on_paste_across_line:
       
  1687   assumes S: "open S" and "d \<noteq> 0"
       
  1688       and holf1: "f holomorphic_on (S \<inter> {z. d \<bullet> z < k})"
       
  1689       and holf2: "f holomorphic_on (S \<inter> {z. k < d \<bullet> z})"
       
  1690       and contf: "continuous_on S f"
       
  1691     shows "f holomorphic_on S"
       
  1692 proof -
       
  1693   have *: "\<exists>t. open t \<and> p \<in> t \<and> continuous_on t f \<and>
       
  1694                (\<forall>a b c. convex hull {a, b, c} \<subseteq> t \<longrightarrow>
       
  1695                          contour_integral (linepath a b) f +
       
  1696                          contour_integral (linepath b c) f +
       
  1697                          contour_integral (linepath c a) f = 0)"
       
  1698           if "p \<in> S" for p
       
  1699   proof -
       
  1700     obtain e where "e>0" and e: "ball p e \<subseteq> S"
       
  1701       using \<open>p \<in> S\<close> openE S by blast
       
  1702     then have "continuous_on (ball p e) f"
       
  1703       using contf continuous_on_subset by blast
       
  1704     moreover have "f holomorphic_on {z. dist p z < e \<and> d \<bullet> z < k}"
       
  1705       apply (rule holomorphic_on_subset [OF holf1])
       
  1706       using e by auto
       
  1707     moreover have "f holomorphic_on {z. dist p z < e \<and> k < d \<bullet> z}"
       
  1708       apply (rule holomorphic_on_subset [OF holf2])
       
  1709       using e by auto
       
  1710     ultimately show ?thesis
       
  1711       apply (rule_tac x="ball p e" in exI)
       
  1712       using \<open>e > 0\<close> e \<open>d \<noteq> 0\<close>
       
  1713       apply (simp add:, clarify)
       
  1714       apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k])
       
  1715       apply (auto simp: subset_hull)
       
  1716       done
       
  1717   qed
       
  1718   show ?thesis
       
  1719     by (blast intro: * Morera_local_triangle analytic_imp_holomorphic)
       
  1720 qed
       
  1721 
       
  1722 proposition Schwarz_reflection:
       
  1723   assumes "open S" and cnjs: "cnj ` S \<subseteq> S"
       
  1724       and  holf: "f holomorphic_on (S \<inter> {z. 0 < Im z})"
       
  1725       and contf: "continuous_on (S \<inter> {z. 0 \<le> Im z}) f"
       
  1726       and f: "\<And>z. \<lbrakk>z \<in> S; z \<in> \<real>\<rbrakk> \<Longrightarrow> (f z) \<in> \<real>"
       
  1727     shows "(\<lambda>z. if 0 \<le> Im z then f z else cnj(f(cnj z))) holomorphic_on S"
       
  1728 proof -
       
  1729   have 1: "(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z))) holomorphic_on (S \<inter> {z. 0 < Im z})"
       
  1730     by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf])
       
  1731   have cont_cfc: "continuous_on (S \<inter> {z. Im z \<le> 0}) (cnj o f o cnj)"
       
  1732     apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf])
       
  1733     using cnjs apply auto
       
  1734     done
       
  1735   have "cnj \<circ> f \<circ> cnj field_differentiable at x within S \<inter> {z. Im z < 0}"
       
  1736         if "x \<in> S" "Im x < 0" "f field_differentiable at (cnj x) within S \<inter> {z. 0 < Im z}" for x
       
  1737     using that
       
  1738     apply (simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm, clarify)
       
  1739     apply (rule_tac x="cnj f'" in exI)
       
  1740     apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify)
       
  1741     apply (drule_tac x="cnj xa" in bspec)
       
  1742     using cnjs apply force
       
  1743     apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj)
       
  1744     done
       
  1745   then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \<inter> {z. Im z < 0})"
       
  1746     using holf cnjs
       
  1747     by (force simp: holomorphic_on_def)
       
  1748   have 2: "(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z))) holomorphic_on (S \<inter> {z. Im z < 0})"
       
  1749     apply (rule iffD1 [OF holomorphic_cong [OF refl]])
       
  1750     using hol_cfc by auto
       
  1751   have [simp]: "(S \<inter> {z. 0 \<le> Im z}) \<union> (S \<inter> {z. Im z \<le> 0}) = S"
       
  1752     by force
       
  1753   have "continuous_on ((S \<inter> {z. 0 \<le> Im z}) \<union> (S \<inter> {z. Im z \<le> 0}))
       
  1754                        (\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z)))"
       
  1755     apply (rule continuous_on_cases_local)
       
  1756     using cont_cfc contf
       
  1757     apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge)
       
  1758     using f Reals_cnj_iff complex_is_Real_iff apply auto
       
  1759     done
       
  1760   then have 3: "continuous_on S (\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z)))"
       
  1761     by force
       
  1762   show ?thesis
       
  1763     apply (rule holomorphic_on_paste_across_line [OF \<open>open S\<close>, of "- \<i>" _ 0])
       
  1764     using 1 2 3
       
  1765     apply auto
       
  1766     done
       
  1767 qed
       
  1768 
       
  1769 subsection\<open>Bloch's theorem\<close>
       
  1770 
       
  1771 lemma Bloch_lemma_0:
       
  1772   assumes holf: "f holomorphic_on cball 0 r" and "0 < r"
       
  1773       and [simp]: "f 0 = 0"
       
  1774       and le: "\<And>z. norm z < r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f 0)"
       
  1775     shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \<subseteq> f ` ball 0 r"
       
  1776 proof -
       
  1777   have "sqrt 2 < 3/2"
       
  1778     by (rule real_less_lsqrt) (auto simp: power2_eq_square)
       
  1779   then have sq3: "0 < 3 - 2 * sqrt 2" by simp
       
  1780   show ?thesis
       
  1781   proof (cases "deriv f 0 = 0")
       
  1782     case True then show ?thesis by simp
       
  1783   next
       
  1784     case False
       
  1785     define C where "C = 2 * norm(deriv f 0)"
       
  1786     have "0 < C" using False by (simp add: C_def)
       
  1787     have holf': "f holomorphic_on ball 0 r" using holf
       
  1788       using ball_subset_cball holomorphic_on_subset by blast
       
  1789     then have holdf': "deriv f holomorphic_on ball 0 r"
       
  1790       by (rule holomorphic_deriv [OF _ open_ball])
       
  1791     have "Le1": "norm(deriv f z - deriv f 0) \<le> norm z / (r - norm z) * C"
       
  1792                 if "norm z < r" for z
       
  1793     proof -
       
  1794       have T1: "norm(deriv f z - deriv f 0) \<le> norm z / (R - norm z) * C"
       
  1795               if R: "norm z < R" "R < r" for R
       
  1796       proof -
       
  1797         have "0 < R" using R
       
  1798           by (metis less_trans norm_zero zero_less_norm_iff)
       
  1799         have df_le: "\<And>x. norm x < r \<Longrightarrow> norm (deriv f x) \<le> C"
       
  1800           using le by (simp add: C_def)
       
  1801         have hol_df: "deriv f holomorphic_on cball 0 R"
       
  1802           apply (rule holomorphic_on_subset) using R holdf' by auto
       
  1803         have *: "((\<lambda>w. deriv f w / (w - z)) has_contour_integral 2 * pi * \<i> * deriv f z) (circlepath 0 R)"
       
  1804                  if "norm z < R" for z
       
  1805           using \<open>0 < R\<close> that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"]
       
  1806           by (force simp: winding_number_circlepath)
       
  1807         have **: "((\<lambda>x. deriv f x / (x - z) - deriv f x / x) has_contour_integral
       
  1808                    of_real (2 * pi) * \<i> * (deriv f z - deriv f 0))
       
  1809                   (circlepath 0 R)"
       
  1810            using has_contour_integral_diff [OF * [of z] * [of 0]] \<open>0 < R\<close> that
       
  1811            by (simp add: algebra_simps)
       
  1812         have [simp]: "\<And>x. norm x = R \<Longrightarrow> x \<noteq> z"  using that(1) by blast
       
  1813         have "norm (deriv f x / (x - z) - deriv f x / x)
       
  1814                      \<le> C * norm z / (R * (R - norm z))"
       
  1815                   if "norm x = R" for x
       
  1816         proof -
       
  1817           have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) =
       
  1818                         norm (deriv f x) * norm z"
       
  1819             by (simp add: norm_mult right_diff_distrib')
       
  1820           show ?thesis
       
  1821             using  \<open>0 < R\<close> \<open>0 < C\<close> R that
       
  1822             apply (simp add: norm_mult norm_divide divide_simps)
       
  1823             using df_le norm_triangle_ineq2 \<open>0 < C\<close> apply (auto intro!: mult_mono)
       
  1824             done
       
  1825         qed
       
  1826         then show ?thesis
       
  1827           using has_contour_integral_bound_circlepath
       
  1828                   [OF **, of "C * norm z/(R*(R - norm z))"]
       
  1829                 \<open>0 < R\<close> \<open>0 < C\<close> R
       
  1830           apply (simp add: norm_mult norm_divide)
       
  1831           apply (simp add: divide_simps mult.commute)
       
  1832           done
       
  1833       qed
       
  1834       obtain r' where r': "norm z < r'" "r' < r"
       
  1835         using Rats_dense_in_real [of "norm z" r] \<open>norm z < r\<close> by blast
       
  1836       then have [simp]: "closure {r'<..<r} = {r'..r}" by simp
       
  1837       show ?thesis
       
  1838         apply (rule continuous_ge_on_closure
       
  1839                  [where f = "\<lambda>r. norm z / (r - norm z) * C" and s = "{r'<..<r}",
       
  1840                   OF _ _ T1])
       
  1841         apply (intro continuous_intros)
       
  1842         using that r'
       
  1843         apply (auto simp: not_le)
       
  1844         done
       
  1845     qed
       
  1846     have "*": "(norm z - norm z^2/(r - norm z)) * norm(deriv f 0) \<le> norm(f z)"
       
  1847               if r: "norm z < r" for z
       
  1848     proof -
       
  1849       have 1: "\<And>x. x \<in> ball 0 r \<Longrightarrow>
       
  1850               ((\<lambda>z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0)
       
  1851                (at x within ball 0 r)"
       
  1852         by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+
       
  1853       have 2: "closed_segment 0 z \<subseteq> ball 0 r"
       
  1854         by (metis \<open>0 < r\<close> convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that)
       
  1855       have 3: "(\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}"
       
  1856         apply (rule integrable_on_cmult_right [where 'b=real, simplified])
       
  1857         apply (rule integrable_on_cdivide [where 'b=real, simplified])
       
  1858         apply (rule integrable_on_cmult_left [where 'b=real, simplified])
       
  1859         apply (rule ident_integrable_on)
       
  1860         done
       
  1861       have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm z * norm z * x * C / (r - norm z)"
       
  1862               if x: "0 \<le> x" "x \<le> 1" for x
       
  1863       proof -
       
  1864         have [simp]: "x * norm z < r"
       
  1865           using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero)
       
  1866         have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \<le> norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C"
       
  1867           apply (rule Le1) using r x \<open>0 < r\<close> by simp
       
  1868         also have "... \<le> norm (x *\<^sub>R z) / (r - norm z) * C"
       
  1869           using r x \<open>0 < r\<close>
       
  1870           apply (simp add: field_split_simps)
       
  1871           by (simp add: \<open>0 < C\<close> mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono)
       
  1872         finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm (x *\<^sub>R z)  / (r - norm z) * C * norm z"
       
  1873           by (rule mult_right_mono) simp
       
  1874         with x show ?thesis by (simp add: algebra_simps)
       
  1875       qed
       
  1876       have le_norm: "abc \<le> norm d - e \<Longrightarrow> norm(f - d) \<le> e \<Longrightarrow> abc \<le> norm f" for abc d e and f::complex
       
  1877         by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans)
       
  1878       have "norm (integral {0..1} (\<lambda>x. (deriv f (x *\<^sub>R z) - deriv f 0) * z))
       
  1879             \<le> integral {0..1} (\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C)"
       
  1880         apply (rule integral_norm_bound_integral)
       
  1881         using contour_integral_primitive [OF 1, of "linepath 0 z"] 2
       
  1882         apply (simp add: has_contour_integral_linepath has_integral_integrable_integral)
       
  1883         apply (rule 3)
       
  1884         apply (simp add: norm_mult power2_eq_square 4)
       
  1885         done
       
  1886       then have int_le: "norm (f z - deriv f 0 * z) \<le> (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))"
       
  1887         using contour_integral_primitive [OF 1, of "linepath 0 z"] 2
       
  1888         apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def)
       
  1889         done
       
  1890       show ?thesis
       
  1891         apply (rule le_norm [OF _ int_le])
       
  1892         using \<open>norm z < r\<close>
       
  1893         apply (simp add: power2_eq_square divide_simps C_def norm_mult)
       
  1894         proof -
       
  1895           have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \<le> norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)"
       
  1896             by (simp add: algebra_simps)
       
  1897           then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \<le> norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)"
       
  1898             by (simp add: algebra_simps)
       
  1899         qed
       
  1900     qed
       
  1901     have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2)  < 1"
       
  1902       by (auto simp:  sqrt2_less_2)
       
  1903     have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f"
       
  1904       apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]])
       
  1905       apply (subst closure_ball)
       
  1906       using \<open>0 < r\<close> mult_pos_pos sq201
       
  1907       apply (auto simp: cball_subset_cball_iff)
       
  1908       done
       
  1909     have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))"
       
  1910       apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force)
       
  1911       using \<open>0 < r\<close> mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff)
       
  1912       using False \<open>0 < r\<close> centre_in_ball holf' holomorphic_nonconstant by blast
       
  1913     have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) =
       
  1914           ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))"
       
  1915       by simp
       
  1916     also have "...  \<subseteq> f ` ball 0 ((1 - sqrt 2 / 2) * r)"
       
  1917     proof -
       
  1918       have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \<le> norm (f z)"
       
  1919            if "norm z = (1 - sqrt 2 / 2) * r" for z
       
  1920         apply (rule order_trans [OF _ *])
       
  1921         using  \<open>0 < r\<close>
       
  1922         apply (simp_all add: field_simps  power2_eq_square that)
       
  1923         apply (simp add: mult.assoc [symmetric])
       
  1924         done
       
  1925       show ?thesis
       
  1926         apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball])
       
  1927         using \<open>0 < r\<close> sq201 3 apply simp_all
       
  1928         using C_def \<open>0 < C\<close> sq3 apply force
       
  1929         done
       
  1930      qed
       
  1931     also have "...  \<subseteq> f ` ball 0 r"
       
  1932       apply (rule image_subsetI [OF imageI], simp)
       
  1933       apply (erule less_le_trans)
       
  1934       using \<open>0 < r\<close> apply (auto simp: field_simps)
       
  1935       done
       
  1936     finally show ?thesis .
       
  1937   qed
       
  1938 qed
       
  1939 
       
  1940 lemma Bloch_lemma:
       
  1941   assumes holf: "f holomorphic_on cball a r" and "0 < r"
       
  1942       and le: "\<And>z. z \<in> ball a r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f a)"
       
  1943     shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \<subseteq> f ` ball a r"
       
  1944 proof -
       
  1945   have fz: "(\<lambda>z. f (a + z)) = f o (\<lambda>z. (a + z))"
       
  1946     by (simp add: o_def)
       
  1947   have hol0: "(\<lambda>z. f (a + z)) holomorphic_on cball 0 r"
       
  1948     unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+
       
  1949   then have [simp]: "\<And>x. norm x < r \<Longrightarrow> (\<lambda>z. f (a + z)) field_differentiable at x"
       
  1950     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)
       
  1951   have [simp]: "\<And>z. norm z < r \<Longrightarrow> f field_differentiable at (a + z)"
       
  1952     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)
       
  1953   then have [simp]: "f field_differentiable at a"
       
  1954     by (metis add.comm_neutral \<open>0 < r\<close> norm_eq_zero)
       
  1955   have hol1: "(\<lambda>z. f (a + z) - f a) holomorphic_on cball 0 r"
       
  1956     by (intro holomorphic_intros hol0)
       
  1957   then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\<lambda>z. f (a + z) - f a) 0))
       
  1958              \<subseteq> (\<lambda>z. f (a + z) - f a) ` ball 0 r"
       
  1959     apply (rule Bloch_lemma_0)
       
  1960     apply (simp_all add: \<open>0 < r\<close>)
       
  1961     apply (simp add: fz complex_derivative_chain)
       
  1962     apply (simp add: dist_norm le)
       
  1963     done
       
  1964   then show ?thesis
       
  1965     apply clarify
       
  1966     apply (drule_tac c="x - f a" in subsetD)
       
  1967      apply (force simp: fz \<open>0 < r\<close> dist_norm complex_derivative_chain field_differentiable_compose)+
       
  1968     done
       
  1969 qed
       
  1970 
       
  1971 proposition Bloch_unit:
       
  1972   assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1"
       
  1973   obtains b r where "1/12 < r" and "ball b r \<subseteq> f ` (ball a 1)"
       
  1974 proof -
       
  1975   define r :: real where "r = 249/256"
       
  1976   have "0 < r" "r < 1" by (auto simp: r_def)
       
  1977   define g where "g z = deriv f z * of_real(r - norm(z - a))" for z
       
  1978   have "deriv f holomorphic_on ball a 1"
       
  1979     by (rule holomorphic_deriv [OF holf open_ball])
       
  1980   then have "continuous_on (ball a 1) (deriv f)"
       
  1981     using holomorphic_on_imp_continuous_on by blast
       
  1982   then have "continuous_on (cball a r) (deriv f)"
       
  1983     by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \<open>r < 1\<close>)
       
  1984   then have "continuous_on (cball a r) g"
       
  1985     by (simp add: g_def continuous_intros)
       
  1986   then have 1: "compact (g ` cball a r)"
       
  1987     by (rule compact_continuous_image [OF _ compact_cball])
       
  1988   have 2: "g ` cball a r \<noteq> {}"
       
  1989     using \<open>r > 0\<close> by auto
       
  1990   obtain p where pr: "p \<in> cball a r"
       
  1991              and pge: "\<And>y. y \<in> cball a r \<Longrightarrow> norm (g y) \<le> norm (g p)"
       
  1992     using distance_attains_sup [OF 1 2, of 0] by force
       
  1993   define t where "t = (r - norm(p - a)) / 2"
       
  1994   have "norm (p - a) \<noteq> r"
       
  1995     using pge [of a] \<open>r > 0\<close> by (auto simp: g_def norm_mult)
       
  1996   then have "norm (p - a) < r" using pr
       
  1997     by (simp add: norm_minus_commute dist_norm)
       
  1998   then have "0 < t"
       
  1999     by (simp add: t_def)
       
  2000   have cpt: "cball p t \<subseteq> ball a r"
       
  2001     using \<open>0 < t\<close> by (simp add: cball_subset_ball_iff dist_norm t_def field_simps)
       
  2002   have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \<le> norm (deriv f p)"
       
  2003             if "y \<in> cball a r" for y
       
  2004   proof -
       
  2005     have [simp]: "norm (y - a) \<le> r"
       
  2006       using that by (simp add: dist_norm norm_minus_commute)
       
  2007     have "norm (g y) \<le> norm (g p)"
       
  2008       using pge [OF that] by simp
       
  2009     then have "norm (deriv f y) * abs (r - norm (y - a)) \<le> norm (deriv f p) * abs (r - norm (p - a))"
       
  2010       by (simp only: dist_norm g_def norm_mult norm_of_real)
       
  2011     with that \<open>norm (p - a) < r\<close> show ?thesis
       
  2012       by (simp add: dist_norm field_split_simps)
       
  2013   qed
       
  2014   have le_norm_dfp: "r / (r - norm (p - a)) \<le> norm (deriv f p)"
       
  2015     using gen_le_dfp [of a] \<open>r > 0\<close> by auto
       
  2016   have 1: "f holomorphic_on cball p t"
       
  2017     apply (rule holomorphic_on_subset [OF holf])
       
  2018     using cpt \<open>r < 1\<close> order_subst1 subset_ball by auto
       
  2019   have 2: "norm (deriv f z) \<le> 2 * norm (deriv f p)" if "z \<in> ball p t" for z
       
  2020   proof -
       
  2021     have z: "z \<in> cball a r"
       
  2022       by (meson ball_subset_cball subsetD cpt that)
       
  2023     then have "norm(z - a) < r"
       
  2024       by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that)
       
  2025     have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \<le> norm (deriv f p)"
       
  2026       using gen_le_dfp [OF z] by simp
       
  2027     with \<open>norm (z - a) < r\<close> \<open>norm (p - a) < r\<close>
       
  2028     have "norm (deriv f z) \<le> (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)"
       
  2029        by (simp add: field_simps)
       
  2030     also have "... \<le> 2 * norm (deriv f p)"
       
  2031       apply (rule mult_right_mono)
       
  2032       using that \<open>norm (p - a) < r\<close> \<open>norm(z - a) < r\<close>
       
  2033       apply (simp_all add: field_simps t_def dist_norm [symmetric])
       
  2034       using dist_triangle3 [of z a p] by linarith
       
  2035     finally show ?thesis .
       
  2036   qed
       
  2037   have sqrt2: "sqrt 2 < 2113/1494"
       
  2038     by (rule real_less_lsqrt) (auto simp: power2_eq_square)
       
  2039   then have sq3: "0 < 3 - 2 * sqrt 2" by simp
       
  2040   have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r"
       
  2041     using sq3 sqrt2 by (auto simp: field_simps r_def)
       
  2042   also have "... \<le> cmod (deriv f p) * (r - cmod (p - a))"
       
  2043     using \<open>norm (p - a) < r\<close> le_norm_dfp   by (simp add: pos_divide_le_eq)
       
  2044   finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)"
       
  2045     using pos_divide_less_eq half_gt_zero_iff sq3 by blast
       
  2046   then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)"
       
  2047     using sq3 by (simp add: mult.commute t_def)
       
  2048   have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball p t"
       
  2049     by (rule Bloch_lemma [OF 1 \<open>0 < t\<close> 2])
       
  2050   also have "... \<subseteq> f ` ball a 1"
       
  2051     apply (rule image_mono)
       
  2052     apply (rule order_trans [OF ball_subset_cball])
       
  2053     apply (rule order_trans [OF cpt])
       
  2054     using \<open>0 < t\<close> \<open>r < 1\<close> apply (simp add: ball_subset_ball_iff dist_norm)
       
  2055     done
       
  2056   finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball a 1" .
       
  2057   with ** show ?thesis
       
  2058     by (rule that)
       
  2059 qed
       
  2060 
       
  2061 theorem Bloch:
       
  2062   assumes holf: "f holomorphic_on ball a r" and "0 < r"
       
  2063       and r': "r' \<le> r * norm (deriv f a) / 12"
       
  2064   obtains b where "ball b r' \<subseteq> f ` (ball a r)"
       
  2065 proof (cases "deriv f a = 0")
       
  2066   case True with r' show ?thesis
       
  2067     using ball_eq_empty that by fastforce
       
  2068 next
       
  2069   case False
       
  2070   define C where "C = deriv f a"
       
  2071   have "0 < norm C" using False by (simp add: C_def)
       
  2072   have dfa: "f field_differentiable at a"
       
  2073     apply (rule holomorphic_on_imp_differentiable_at [OF holf])
       
  2074     using \<open>0 < r\<close> by auto
       
  2075   have fo: "(\<lambda>z. f (a + of_real r * z)) = f o (\<lambda>z. (a + of_real r * z))"
       
  2076     by (simp add: o_def)
       
  2077   have holf': "f holomorphic_on (\<lambda>z. a + complex_of_real r * z) ` ball 0 1"
       
  2078     apply (rule holomorphic_on_subset [OF holf])
       
  2079     using \<open>0 < r\<close> apply (force simp: dist_norm norm_mult)
       
  2080     done
       
  2081   have 1: "(\<lambda>z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1"
       
  2082     apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+
       
  2083     using \<open>0 < r\<close> by (simp add: C_def False)
       
  2084   have "((\<lambda>z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative
       
  2085         (deriv f (a + of_real r * z) / C)) (at z)"
       
  2086        if "norm z < 1" for z
       
  2087   proof -
       
  2088     have *: "((\<lambda>x. f (a + of_real r * x)) has_field_derivative
       
  2089            (deriv f (a + of_real r * z) * of_real r)) (at z)"
       
  2090       apply (simp add: fo)
       
  2091       apply (rule DERIV_chain [OF field_differentiable_derivI])
       
  2092       apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp)
       
  2093       using \<open>0 < r\<close> apply (simp add: dist_norm norm_mult that)
       
  2094       apply (rule derivative_eq_intros | simp)+
       
  2095       done
       
  2096     show ?thesis
       
  2097       apply (rule derivative_eq_intros * | simp)+
       
  2098       using \<open>0 < r\<close> by (auto simp: C_def False)
       
  2099   qed
       
  2100   have 2: "deriv (\<lambda>z. f (a + of_real r * z) / (C * of_real r)) 0 = 1"
       
  2101     apply (subst deriv_cdivide_right)
       
  2102     apply (simp add: field_differentiable_def fo)
       
  2103     apply (rule exI)
       
  2104     apply (rule DERIV_chain [OF field_differentiable_derivI])
       
  2105     apply (simp add: dfa)
       
  2106     apply (rule derivative_eq_intros | simp add: C_def False fo)+
       
  2107     using \<open>0 < r\<close>
       
  2108     apply (simp add: C_def False fo)
       
  2109     apply (simp add: derivative_intros dfa complex_derivative_chain)
       
  2110     done
       
  2111   have sb1: "(*) (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1
       
  2112              \<subseteq> f ` ball a r"
       
  2113     using \<open>0 < r\<close> by (auto simp: dist_norm norm_mult C_def False)
       
  2114   have sb2: "ball (C * r * b) r' \<subseteq> (*) (C * r) ` ball b t"
       
  2115              if "1 / 12 < t" for b t
       
  2116   proof -
       
  2117     have *: "r * cmod (deriv f a) / 12 \<le> r * (t * cmod (deriv f a))"
       
  2118       using that \<open>0 < r\<close> less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right
       
  2119       by auto
       
  2120     show ?thesis
       
  2121       apply clarify
       
  2122       apply (rule_tac x="x / (C * r)" in image_eqI)
       
  2123       using \<open>0 < r\<close>
       
  2124       apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps)
       
  2125       apply (erule less_le_trans)
       
  2126       apply (rule order_trans [OF r' *])
       
  2127       done
       
  2128   qed
       
  2129   show ?thesis
       
  2130     apply (rule Bloch_unit [OF 1 2])
       
  2131     apply (rename_tac t)
       
  2132     apply (rule_tac b="(C * of_real r) * b" in that)
       
  2133     apply (drule image_mono [where f = "\<lambda>z. (C * of_real r) * z"])
       
  2134     using sb1 sb2
       
  2135     apply force
       
  2136     done
       
  2137 qed
       
  2138 
       
  2139 corollary Bloch_general:
       
  2140   assumes holf: "f holomorphic_on s" and "a \<in> s"
       
  2141       and tle: "\<And>z. z \<in> frontier s \<Longrightarrow> t \<le> dist a z"
       
  2142       and rle: "r \<le> t * norm(deriv f a) / 12"
       
  2143   obtains b where "ball b r \<subseteq> f ` s"
       
  2144 proof -
       
  2145   consider "r \<le> 0" | "0 < t * norm(deriv f a) / 12" using rle by force
       
  2146   then show ?thesis
       
  2147   proof cases
       
  2148     case 1 then show ?thesis
       
  2149       by (simp add: ball_empty that)
       
  2150   next
       
  2151     case 2
       
  2152     show ?thesis
       
  2153     proof (cases "deriv f a = 0")
       
  2154       case True then show ?thesis
       
  2155         using rle by (simp add: ball_empty that)
       
  2156     next
       
  2157       case False
       
  2158       then have "t > 0"
       
  2159         using 2 by (force simp: zero_less_mult_iff)
       
  2160       have "\<not> ball a t \<subseteq> s \<Longrightarrow> ball a t \<inter> frontier s \<noteq> {}"
       
  2161         apply (rule connected_Int_frontier [of "ball a t" s], simp_all)
       
  2162         using \<open>0 < t\<close> \<open>a \<in> s\<close> centre_in_ball apply blast
       
  2163         done
       
  2164       with tle have *: "ball a t \<subseteq> s" by fastforce
       
  2165       then have 1: "f holomorphic_on ball a t"
       
  2166         using holf using holomorphic_on_subset by blast
       
  2167       show ?thesis
       
  2168         apply (rule Bloch [OF 1 \<open>t > 0\<close> rle])
       
  2169         apply (rule_tac b=b in that)
       
  2170         using * apply force
       
  2171         done
       
  2172     qed
       
  2173   qed
       
  2174 qed
       
  2175 
       
  2176 subsection \<open>Cauchy's residue theorem\<close>
       
  2177 
       
  2178 text\<open>Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem.
       
  2179     Interactive Theorem Proving\<close>
       
  2180 
       
  2181 definition\<^marker>\<open>tag important\<close> residue :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where
       
  2182   "residue f z = (SOME int. \<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e
       
  2183     \<longrightarrow> (f has_contour_integral 2*pi* \<i> *int) (circlepath z \<epsilon>))"
       
  2184 
       
  2185 lemma Eps_cong:
       
  2186   assumes "\<And>x. P x = Q x"
       
  2187   shows   "Eps P = Eps Q"
       
  2188   using ext[of P Q, OF assms] by simp
       
  2189 
       
  2190 lemma residue_cong:
       
  2191   assumes eq: "eventually (\<lambda>z. f z = g z) (at z)" and "z = z'"
       
  2192   shows   "residue f z = residue g z'"
       
  2193 proof -
       
  2194   from assms have eq': "eventually (\<lambda>z. g z = f z) (at z)"
       
  2195     by (simp add: eq_commute)
       
  2196   let ?P = "\<lambda>f c e. (\<forall>\<epsilon>>0. \<epsilon> < e \<longrightarrow>
       
  2197    (f has_contour_integral of_real (2 * pi) * \<i> * c) (circlepath z \<epsilon>))"
       
  2198   have "residue f z = residue g z" unfolding residue_def
       
  2199   proof (rule Eps_cong)
       
  2200     fix c :: complex
       
  2201     have "\<exists>e>0. ?P g c e"
       
  2202       if "\<exists>e>0. ?P f c e" and "eventually (\<lambda>z. f z = g z) (at z)" for f g
       
  2203     proof -
       
  2204       from that(1) obtain e where e: "e > 0" "?P f c e"
       
  2205         by blast
       
  2206       from that(2) obtain e' where e': "e' > 0" "\<And>z'. z' \<noteq> z \<Longrightarrow> dist z' z < e' \<Longrightarrow> f z' = g z'"
       
  2207         unfolding eventually_at by blast
       
  2208       have "?P g c (min e e')"
       
  2209       proof (intro allI exI impI, goal_cases)
       
  2210         case (1 \<epsilon>)
       
  2211         hence "(f has_contour_integral of_real (2 * pi) * \<i> * c) (circlepath z \<epsilon>)"
       
  2212           using e(2) by auto
       
  2213         thus ?case
       
  2214         proof (rule has_contour_integral_eq)
       
  2215           fix z' assume "z' \<in> path_image (circlepath z \<epsilon>)"
       
  2216           hence "dist z' z < e'" and "z' \<noteq> z"
       
  2217             using 1 by (auto simp: dist_commute)
       
  2218           with e'(2)[of z'] show "f z' = g z'" by simp
       
  2219         qed
       
  2220       qed
       
  2221       moreover from e and e' have "min e e' > 0" by auto
       
  2222       ultimately show ?thesis by blast
       
  2223     qed
       
  2224     from this[OF _ eq] and this[OF _ eq']
       
  2225       show "(\<exists>e>0. ?P f c e) \<longleftrightarrow> (\<exists>e>0. ?P g c e)"
       
  2226       by blast
       
  2227   qed
       
  2228   with assms show ?thesis by simp
       
  2229 qed
       
  2230 
       
  2231 lemma contour_integral_circlepath_eq:
       
  2232   assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0<e1" "e1\<le>e2"
       
  2233     and e2_cball:"cball z e2 \<subseteq> s"
       
  2234   shows
       
  2235     "f contour_integrable_on circlepath z e1"
       
  2236     "f contour_integrable_on circlepath z e2"
       
  2237     "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f"
       
  2238 proof -
       
  2239   define l where "l \<equiv> linepath (z+e2) (z+e1)"
       
  2240   have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto
       
  2241   have "e2>0" using \<open>e1>0\<close> \<open>e1\<le>e2\<close> by auto
       
  2242   have zl_img:"z\<notin>path_image l"
       
  2243     proof
       
  2244       assume "z \<in> path_image l"
       
  2245       then have "e2 \<le> cmod (e2 - e1)"
       
  2246         using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \<open>e1>0\<close> \<open>e2>0\<close> unfolding l_def
       
  2247         by (auto simp add:closed_segment_commute)
       
  2248       thus False using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close>
       
  2249         apply (subst (asm) norm_of_real)
       
  2250         by auto
       
  2251     qed
       
  2252   define g where "g \<equiv> circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l"
       
  2253   show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)"
       
  2254     proof -
       
  2255       show "f contour_integrable_on circlepath z e2"
       
  2256         apply (intro contour_integrable_continuous_circlepath[OF
       
  2257                 continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
       
  2258         using \<open>e2>0\<close> e2_cball by auto
       
  2259       show "f contour_integrable_on (circlepath z e1)"
       
  2260         apply (intro contour_integrable_continuous_circlepath[OF
       
  2261                       continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
       
  2262         using \<open>e1>0\<close> \<open>e1\<le>e2\<close> e2_cball by auto
       
  2263     qed
       
  2264   have [simp]:"f contour_integrable_on l"
       
  2265     proof -
       
  2266       have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close>
       
  2267         by (intro closed_segment_subset,auto simp add:dist_norm)
       
  2268       hence "closed_segment (z + e2) (z + e1) \<subseteq> s - {z}" using zl_img e2_cball unfolding l_def
       
  2269         by auto
       
  2270       then show "f contour_integrable_on l" unfolding l_def
       
  2271         apply (intro contour_integrable_continuous_linepath[OF
       
  2272                       continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
       
  2273         by auto
       
  2274     qed
       
  2275   let ?ig="\<lambda>g. contour_integral g f"
       
  2276   have "(f has_contour_integral 0) g"
       
  2277     proof (rule Cauchy_theorem_global[OF _ f_holo])
       
  2278       show "open (s - {z})" using \<open>open s\<close> by auto
       
  2279       show "valid_path g" unfolding g_def l_def by auto
       
  2280       show "pathfinish g = pathstart g" unfolding g_def l_def by auto
       
  2281     next
       
  2282       have path_img:"path_image g \<subseteq> cball z e2"
       
  2283         proof -
       
  2284           have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close>
       
  2285             by (intro closed_segment_subset,auto simp add:dist_norm)
       
  2286           moreover have "sphere z \<bar>e1\<bar> \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1\<le>e2\<close> \<open>e1>0\<close> by auto
       
  2287           ultimately show ?thesis unfolding g_def l_def using \<open>e2>0\<close>
       
  2288             by (simp add: path_image_join closed_segment_commute)
       
  2289         qed
       
  2290       show "path_image g \<subseteq> s - {z}"
       
  2291         proof -
       
  2292           have "z\<notin>path_image g" using zl_img
       
  2293             unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute)
       
  2294           moreover note \<open>cball z e2 \<subseteq> s\<close> and path_img
       
  2295           ultimately show ?thesis by auto
       
  2296         qed
       
  2297       show "winding_number g w = 0" when"w \<notin> s - {z}" for w
       
  2298         proof -
       
  2299           have "winding_number g w = 0" when "w\<notin>s" using that e2_cball
       
  2300             apply (intro winding_number_zero_outside[OF _ _ _ _ path_img])
       
  2301             by (auto simp add:g_def l_def)
       
  2302           moreover have "winding_number g z=0"
       
  2303             proof -
       
  2304               let ?Wz="\<lambda>g. winding_number g z"
       
  2305               have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1))
       
  2306                   + ?Wz (reversepath l)"
       
  2307                 using \<open>e2>0\<close> \<open>e1>0\<close> zl_img unfolding g_def l_def
       
  2308                 by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+
       
  2309               also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))"
       
  2310                 using zl_img
       
  2311                 apply (subst (2) winding_number_reversepath)
       
  2312                 by (auto simp add:l_def closed_segment_commute)
       
  2313               also have "... = 0"
       
  2314                 proof -
       
  2315                   have "?Wz (circlepath z e2) = 1" using \<open>e2>0\<close>
       
  2316                     by (auto intro: winding_number_circlepath_centre)
       
  2317                   moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \<open>e1>0\<close>
       
  2318                     apply (subst winding_number_reversepath)
       
  2319                     by (auto intro: winding_number_circlepath_centre)
       
  2320                   ultimately show ?thesis by auto
       
  2321                 qed
       
  2322               finally show ?thesis .
       
  2323             qed
       
  2324           ultimately show ?thesis using that by auto
       
  2325         qed
       
  2326     qed
       
  2327   then have "0 = ?ig g" using contour_integral_unique by simp
       
  2328   also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1))
       
  2329       + ?ig (reversepath l)"
       
  2330     unfolding g_def
       
  2331     by (auto simp add:contour_integrable_reversepath_eq)
       
  2332   also have "... = ?ig (circlepath z e2)  - ?ig (circlepath z e1)"
       
  2333     by (auto simp add:contour_integral_reversepath)
       
  2334   finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f"
       
  2335     by simp
       
  2336 qed
       
  2337 
       
  2338 lemma base_residue:
       
  2339   assumes "open s" "z\<in>s" "r>0" and f_holo:"f holomorphic_on (s - {z})"
       
  2340     and r_cball:"cball z r \<subseteq> s"
       
  2341   shows "(f has_contour_integral 2 * pi * \<i> * (residue f z)) (circlepath z r)"
       
  2342 proof -
       
  2343   obtain e where "e>0" and e_cball:"cball z e \<subseteq> s"
       
  2344     using open_contains_cball[of s] \<open>open s\<close> \<open>z\<in>s\<close> by auto
       
  2345   define c where "c \<equiv> 2 * pi * \<i>"
       
  2346   define i where "i \<equiv> contour_integral (circlepath z e) f / c"
       
  2347   have "(f has_contour_integral c*i) (circlepath z \<epsilon>)" when "\<epsilon>>0" "\<epsilon><e" for \<epsilon>
       
  2348     proof -
       
  2349       have "contour_integral (circlepath z e) f = contour_integral (circlepath z \<epsilon>) f"
       
  2350           "f contour_integrable_on circlepath z \<epsilon>"
       
  2351           "f contour_integrable_on circlepath z e"
       
  2352         using \<open>\<epsilon><e\<close>
       
  2353         by (intro contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> _ e_cball],auto)+
       
  2354       then show ?thesis unfolding i_def c_def
       
  2355         by (auto intro:has_contour_integral_integral)
       
  2356     qed
       
  2357   then have "\<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
       
  2358     unfolding residue_def c_def
       
  2359     apply (rule_tac someI[of _ i],intro  exI[where x=e])
       
  2360     by (auto simp add:\<open>e>0\<close> c_def)
       
  2361   then obtain e' where "e'>0"
       
  2362       and e'_def:"\<forall>\<epsilon>>0. \<epsilon><e' \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
       
  2363     by auto
       
  2364   let ?int="\<lambda>e. contour_integral (circlepath z e) f"
       
  2365   define  \<epsilon> where "\<epsilon> \<equiv> Min {r,e'} / 2"
       
  2366   have "\<epsilon>>0" "\<epsilon>\<le>r" "\<epsilon><e'" using \<open>r>0\<close> \<open>e'>0\<close> unfolding \<epsilon>_def by auto
       
  2367   have "(f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
       
  2368     using e'_def[rule_format,OF \<open>\<epsilon>>0\<close> \<open>\<epsilon><e'\<close>] .
       
  2369   then show ?thesis unfolding c_def
       
  2370     using contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> \<open>\<epsilon>\<le>r\<close> r_cball]
       
  2371     by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \<epsilon>" "circlepath z r"])
       
  2372 qed
       
  2373 
       
  2374 lemma residue_holo:
       
  2375   assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s"
       
  2376   shows "residue f z = 0"
       
  2377 proof -
       
  2378   define c where "c \<equiv> 2 * pi * \<i>"
       
  2379   obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
       
  2380     using open_contains_cball_eq by blast
       
  2381   have "(f has_contour_integral c*residue f z) (circlepath z e)"
       
  2382     using f_holo
       
  2383     by (auto intro: base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
       
  2384   moreover have "(f has_contour_integral 0) (circlepath z e)"
       
  2385     using f_holo e_cball \<open>e>0\<close>
       
  2386     by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"])
       
  2387   ultimately have "c*residue f z =0"
       
  2388     using has_contour_integral_unique by blast
       
  2389   thus ?thesis unfolding c_def  by auto
       
  2390 qed
       
  2391 
       
  2392 lemma residue_const:"residue (\<lambda>_. c) z = 0"
       
  2393   by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros)
       
  2394 
       
  2395 lemma residue_add:
       
  2396   assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
       
  2397       and g_holo:"g holomorphic_on s - {z}"
       
  2398   shows "residue (\<lambda>z. f z + g z) z= residue f z + residue g z"
       
  2399 proof -
       
  2400   define c where "c \<equiv> 2 * pi * \<i>"
       
  2401   define fg where "fg \<equiv> (\<lambda>z. f z+g z)"
       
  2402   obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
       
  2403     using open_contains_cball_eq by blast
       
  2404   have "(fg has_contour_integral c * residue fg z) (circlepath z e)"
       
  2405     unfolding fg_def using f_holo g_holo
       
  2406     apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
       
  2407     by (auto intro:holomorphic_intros)
       
  2408   moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)"
       
  2409     unfolding fg_def using f_holo g_holo
       
  2410     by (auto intro: has_contour_integral_add base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
       
  2411   ultimately have "c*(residue f z + residue g z) = c * residue fg z"
       
  2412     using has_contour_integral_unique by (auto simp add:distrib_left)
       
  2413   thus ?thesis unfolding fg_def
       
  2414     by (auto simp add:c_def)
       
  2415 qed
       
  2416 
       
  2417 lemma residue_lmul:
       
  2418   assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
       
  2419   shows "residue (\<lambda>z. c * (f z)) z= c * residue f z"
       
  2420 proof (cases "c=0")
       
  2421   case True
       
  2422   thus ?thesis using residue_const by auto
       
  2423 next
       
  2424   case False
       
  2425   define c' where "c' \<equiv> 2 * pi * \<i>"
       
  2426   define f' where "f' \<equiv> (\<lambda>z. c * (f z))"
       
  2427   obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
       
  2428     using open_contains_cball_eq by blast
       
  2429   have "(f' has_contour_integral c' * residue f' z) (circlepath z e)"
       
  2430     unfolding f'_def using f_holo
       
  2431     apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def])
       
  2432     by (auto intro:holomorphic_intros)
       
  2433   moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)"
       
  2434     unfolding f'_def using f_holo
       
  2435     by (auto intro: has_contour_integral_lmul
       
  2436       base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def])
       
  2437   ultimately have "c' * residue f' z  = c * (c' * residue f z)"
       
  2438     using has_contour_integral_unique by auto
       
  2439   thus ?thesis unfolding f'_def c'_def using False
       
  2440     by (auto simp add:field_simps)
       
  2441 qed
       
  2442 
       
  2443 lemma residue_rmul:
       
  2444   assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
       
  2445   shows "residue (\<lambda>z. (f z) * c) z= residue f z * c"
       
  2446 using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps)
       
  2447 
       
  2448 lemma residue_div:
       
  2449   assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
       
  2450   shows "residue (\<lambda>z. (f z) / c) z= residue f z / c "
       
  2451 using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps)
       
  2452 
       
  2453 lemma residue_neg:
       
  2454   assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
       
  2455   shows "residue (\<lambda>z. - (f z)) z= - residue f z"
       
  2456 using residue_lmul[OF assms,of "-1"] by auto
       
  2457 
       
  2458 lemma residue_diff:
       
  2459   assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
       
  2460       and g_holo:"g holomorphic_on s - {z}"
       
  2461   shows "residue (\<lambda>z. f z - g z) z= residue f z - residue g z"
       
  2462 using residue_add[OF assms(1,2,3),of "\<lambda>z. - g z"] residue_neg[OF assms(1,2,4)]
       
  2463 by (auto intro:holomorphic_intros g_holo)
       
  2464 
       
  2465 lemma residue_simple:
       
  2466   assumes "open s" "z\<in>s" and f_holo:"f holomorphic_on s"
       
  2467   shows "residue (\<lambda>w. f w / (w - z)) z = f z"
       
  2468 proof -
       
  2469   define c where "c \<equiv> 2 * pi * \<i>"
       
  2470   define f' where "f' \<equiv> \<lambda>w. f w / (w - z)"
       
  2471   obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
       
  2472     using open_contains_cball_eq by blast
       
  2473   have "(f' has_contour_integral c * f z) (circlepath z e)"
       
  2474     unfolding f'_def c_def using \<open>e>0\<close> f_holo e_cball
       
  2475     by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros)
       
  2476   moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)"
       
  2477     unfolding f'_def using f_holo
       
  2478     apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
       
  2479     by (auto intro!:holomorphic_intros)
       
  2480   ultimately have "c * f z = c * residue f' z"
       
  2481     using has_contour_integral_unique by blast
       
  2482   thus ?thesis unfolding c_def f'_def  by auto
       
  2483 qed
       
  2484 
       
  2485 lemma residue_simple':
       
  2486   assumes s: "open s" "z \<in> s" and holo: "f holomorphic_on (s - {z})"
       
  2487       and lim: "((\<lambda>w. f w * (w - z)) \<longlongrightarrow> c) (at z)"
       
  2488   shows   "residue f z = c"
       
  2489 proof -
       
  2490   define g where "g = (\<lambda>w. if w = z then c else f w * (w - z))"
       
  2491   from holo have "(\<lambda>w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P")
       
  2492     by (force intro: holomorphic_intros)
       
  2493   also have "?P \<longleftrightarrow> g holomorphic_on (s - {z})"
       
  2494     by (intro holomorphic_cong refl) (simp_all add: g_def)
       
  2495   finally have *: "g holomorphic_on (s - {z})" .
       
  2496 
       
  2497   note lim
       
  2498   also have "(\<lambda>w. f w * (w - z)) \<midarrow>z\<rightarrow> c \<longleftrightarrow> g \<midarrow>z\<rightarrow> g z"
       
  2499     by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter)
       
  2500   finally have **: "g \<midarrow>z\<rightarrow> g z" .
       
  2501 
       
  2502   have g_holo: "g holomorphic_on s"
       
  2503     by (rule no_isolated_singularity'[where K = "{z}"])
       
  2504        (insert assms * **, simp_all add: at_within_open_NO_MATCH)
       
  2505   from s and this have "residue (\<lambda>w. g w / (w - z)) z = g z"
       
  2506     by (rule residue_simple)
       
  2507   also have "\<forall>\<^sub>F za in at z. g za / (za - z) = f za"
       
  2508     unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def)
       
  2509   hence "residue (\<lambda>w. g w / (w - z)) z = residue f z"
       
  2510     by (intro residue_cong refl)
       
  2511   finally show ?thesis
       
  2512     by (simp add: g_def)
       
  2513 qed
       
  2514 
       
  2515 lemma residue_holomorphic_over_power:
       
  2516   assumes "open A" "z0 \<in> A" "f holomorphic_on A"
       
  2517   shows   "residue (\<lambda>z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n"
       
  2518 proof -
       
  2519   let ?f = "\<lambda>z. f z / (z - z0) ^ Suc n"
       
  2520   from assms(1,2) obtain r where r: "r > 0" "cball z0 r \<subseteq> A"
       
  2521     by (auto simp: open_contains_cball)
       
  2522   have "(?f has_contour_integral 2 * pi * \<i> * residue ?f z0) (circlepath z0 r)"
       
  2523     using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros)
       
  2524   moreover have "(?f has_contour_integral 2 * pi * \<i> / fact n * (deriv ^^ n) f z0) (circlepath z0 r)"
       
  2525     using assms r
       
  2526     by (intro Cauchy_has_contour_integral_higher_derivative_circlepath)
       
  2527        (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on)
       
  2528   ultimately have "2 * pi * \<i> * residue ?f z0 = 2 * pi * \<i> / fact n * (deriv ^^ n) f z0"
       
  2529     by (rule has_contour_integral_unique)
       
  2530   thus ?thesis by (simp add: field_simps)
       
  2531 qed
       
  2532 
       
  2533 lemma residue_holomorphic_over_power':
       
  2534   assumes "open A" "0 \<in> A" "f holomorphic_on A"
       
  2535   shows   "residue (\<lambda>z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n"
       
  2536   using residue_holomorphic_over_power[OF assms] by simp
       
  2537 
       
  2538 lemma get_integrable_path:
       
  2539   assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\<in>s-pts" "b\<in>s-pts"
       
  2540   obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b"
       
  2541     "path_image g \<subseteq> s-pts" "f contour_integrable_on g" using assms
       
  2542 proof (induct arbitrary:s thesis a rule:finite_induct[OF \<open>finite pts\<close>])
       
  2543   case 1
       
  2544   obtain g where "valid_path g" "path_image g \<subseteq> s" "pathstart g = a" "pathfinish g = b"
       
  2545     using connected_open_polynomial_connected[OF \<open>open s\<close>,of a b ] \<open>connected (s - {})\<close>
       
  2546       valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto
       
  2547   moreover have "f contour_integrable_on g"
       
  2548     using contour_integrable_holomorphic_simple[OF _ \<open>open s\<close> \<open>valid_path g\<close> \<open>path_image g \<subseteq> s\<close>,of f]
       
  2549       \<open>f holomorphic_on s - {}\<close>
       
  2550     by auto
       
  2551   ultimately show ?case using "1"(1)[of g] by auto
       
  2552 next
       
  2553   case idt:(2 p pts)
       
  2554   obtain e where "e>0" and e:"\<forall>w\<in>ball a e. w \<in> s \<and> (w \<noteq> a \<longrightarrow> w \<notin> insert p pts)"
       
  2555     using finite_ball_avoid[OF \<open>open s\<close> \<open>finite (insert p pts)\<close>, of a]
       
  2556       \<open>a \<in> s - insert p pts\<close>
       
  2557     by auto
       
  2558   define a' where "a' \<equiv> a+e/2"
       
  2559   have "a'\<in>s-{p} -pts"  using e[rule_format,of "a+e/2"] \<open>e>0\<close>
       
  2560     by (auto simp add:dist_complex_def a'_def)
       
  2561   then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b"
       
  2562     "path_image g' \<subseteq> s - {p} - pts" "f contour_integrable_on g'"
       
  2563     using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1)
       
  2564     by (metis Diff_insert2 open_delete)
       
  2565   define g where "g \<equiv> linepath a a' +++ g'"
       
  2566   have "valid_path g" unfolding g_def by (auto intro: valid_path_join)
       
  2567   moreover have "pathstart g = a" and  "pathfinish g = b" unfolding g_def by auto
       
  2568   moreover have "path_image g \<subseteq> s - insert p pts" unfolding g_def
       
  2569     proof (rule subset_path_image_join)
       
  2570       have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
       
  2571         by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
       
  2572       then show "path_image (linepath a a') \<subseteq> s - insert p pts" using e idt(9)
       
  2573         by auto
       
  2574     next
       
  2575       show "path_image g' \<subseteq> s - insert p pts" using g'(4) by blast
       
  2576     qed
       
  2577   moreover have "f contour_integrable_on g"
       
  2578     proof -
       
  2579       have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
       
  2580         by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
       
  2581       then have "continuous_on (closed_segment a a') f"
       
  2582         using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)]
       
  2583         apply (elim continuous_on_subset)
       
  2584         by auto
       
  2585       then have "f contour_integrable_on linepath a a'"
       
  2586         using contour_integrable_continuous_linepath by auto
       
  2587       then show ?thesis unfolding g_def
       
  2588         apply (rule contour_integrable_joinI)
       
  2589         by (auto simp add: \<open>e>0\<close>)
       
  2590     qed
       
  2591   ultimately show ?case using idt.prems(1)[of g] by auto
       
  2592 qed
       
  2593 
       
  2594 lemma Cauchy_theorem_aux:
       
  2595   assumes "open s" "connected (s-pts)" "finite pts" "pts \<subseteq> s" "f holomorphic_on s-pts"
       
  2596           "valid_path g" "pathfinish g = pathstart g" "path_image g \<subseteq> s-pts"
       
  2597           "\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0"
       
  2598           "\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
       
  2599   shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
       
  2600     using assms
       
  2601 proof (induct arbitrary:s g rule:finite_induct[OF \<open>finite pts\<close>])
       
  2602   case 1
       
  2603   then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique)
       
  2604 next
       
  2605   case (2 p pts)
       
  2606   note fin[simp] = \<open>finite (insert p pts)\<close>
       
  2607     and connected = \<open>connected (s - insert p pts)\<close>
       
  2608     and valid[simp] = \<open>valid_path g\<close>
       
  2609     and g_loop[simp] = \<open>pathfinish g = pathstart g\<close>
       
  2610     and holo[simp]= \<open>f holomorphic_on s - insert p pts\<close>
       
  2611     and path_img = \<open>path_image g \<subseteq> s - insert p pts\<close>
       
  2612     and winding = \<open>\<forall>z. z \<notin> s \<longrightarrow> winding_number g z = 0\<close>
       
  2613     and h = \<open>\<forall>pa\<in>s. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s \<and> (w \<noteq> pa \<longrightarrow> w \<notin> insert p pts))\<close>
       
  2614   have "h p>0" and "p\<in>s"
       
  2615     and h_p: "\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> insert p pts)"
       
  2616     using h \<open>insert p pts \<subseteq> s\<close> by auto
       
  2617   obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p"
       
  2618       "path_image pg \<subseteq> s-insert p pts" "f contour_integrable_on pg"
       
  2619     proof -
       
  2620       have "p + h p\<in>cball p (h p)" using h[rule_format,of p]
       
  2621         by (simp add: \<open>p \<in> s\<close> dist_norm)
       
  2622       then have "p + h p \<in> s - insert p pts" using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close>
       
  2623         by fastforce
       
  2624       moreover have "pathstart g \<in> s - insert p pts " using path_img by auto
       
  2625       ultimately show ?thesis
       
  2626         using get_integrable_path[OF \<open>open s\<close> connected fin holo,of "pathstart g" "p+h p"] that
       
  2627         by blast
       
  2628     qed
       
  2629   obtain n::int where "n=winding_number g p"
       
  2630     using integer_winding_number[OF _ g_loop,of p] valid path_img
       
  2631     by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path)
       
  2632   define p_circ where "p_circ \<equiv> circlepath p (h p)"
       
  2633   define p_circ_pt where "p_circ_pt \<equiv> linepath (p+h p) (p+h p)"
       
  2634   define n_circ where "n_circ \<equiv> \<lambda>n. ((+++) p_circ ^^ n) p_circ_pt"
       
  2635   define cp where "cp \<equiv> if n\<ge>0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))"
       
  2636   have n_circ:"valid_path (n_circ k)"
       
  2637       "winding_number (n_circ k) p = k"
       
  2638       "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p"
       
  2639       "path_image (n_circ k) =  (if k=0 then {p + h p} else sphere p (h p))"
       
  2640       "p \<notin> path_image (n_circ k)"
       
  2641       "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number (n_circ k) p'=0 \<and> p'\<notin>path_image (n_circ k)"
       
  2642       "f contour_integrable_on (n_circ k)"
       
  2643       "contour_integral (n_circ k) f = k *  contour_integral p_circ f"
       
  2644       for k
       
  2645     proof (induct k)
       
  2646       case 0
       
  2647       show "valid_path (n_circ 0)"
       
  2648         and "path_image (n_circ 0) =  (if 0=0 then {p + h p} else sphere p (h p))"
       
  2649         and "winding_number (n_circ 0) p = of_nat 0"
       
  2650         and "pathstart (n_circ 0) = p + h p"
       
  2651         and "pathfinish (n_circ 0) = p + h p"
       
  2652         and "p \<notin> path_image (n_circ 0)"
       
  2653         unfolding n_circ_def p_circ_pt_def using \<open>h p > 0\<close>
       
  2654         by (auto simp add: dist_norm)
       
  2655       show "winding_number (n_circ 0) p'=0 \<and> p'\<notin>path_image (n_circ 0)" when "p'\<notin>s- pts" for p'
       
  2656         unfolding n_circ_def p_circ_pt_def
       
  2657         apply (auto intro!:winding_number_trivial)
       
  2658         by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+
       
  2659       show "f contour_integrable_on (n_circ 0)"
       
  2660         unfolding n_circ_def p_circ_pt_def
       
  2661         by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing)
       
  2662       show "contour_integral (n_circ 0) f = of_nat 0  *  contour_integral p_circ f"
       
  2663         unfolding n_circ_def p_circ_pt_def by auto
       
  2664     next
       
  2665       case (Suc k)
       
  2666       have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto
       
  2667       have pcirc:"p \<notin> path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)"
       
  2668         using Suc(3) unfolding p_circ_def using \<open>h p > 0\<close> by (auto simp add: p_circ_def)
       
  2669       have pcirc_image:"path_image p_circ \<subseteq> s - insert p pts"
       
  2670         proof -
       
  2671           have "path_image p_circ \<subseteq> cball p (h p)" using \<open>0 < h p\<close> p_circ_def by auto
       
  2672           then show ?thesis using h_p pcirc(1) by auto
       
  2673         qed
       
  2674       have pcirc_integrable:"f contour_integrable_on p_circ"
       
  2675         by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def]
       
  2676           contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on
       
  2677           holomorphic_on_subset[OF holo])
       
  2678       show "valid_path (n_circ (Suc k))"
       
  2679         using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto
       
  2680       show "path_image (n_circ (Suc k))
       
  2681           = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))"
       
  2682         proof -
       
  2683           have "path_image p_circ = sphere p (h p)"
       
  2684             unfolding p_circ_def using \<open>0 < h p\<close> by auto
       
  2685           then show ?thesis unfolding n_Suc  using Suc.hyps(5)  \<open>h p>0\<close>
       
  2686             by (auto simp add:  path_image_join[OF pcirc(3)]  dist_norm)
       
  2687         qed
       
  2688       then show "p \<notin> path_image (n_circ (Suc k))" using \<open>h p>0\<close> by auto
       
  2689       show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)"
       
  2690         proof -
       
  2691           have "winding_number p_circ p = 1"
       
  2692             by (simp add: \<open>h p > 0\<close> p_circ_def winding_number_circlepath_centre)
       
  2693           moreover have "p \<notin> path_image (n_circ k)" using Suc(5) \<open>h p>0\<close> by auto
       
  2694           then have "winding_number (p_circ +++ n_circ k) p
       
  2695               = winding_number p_circ p + winding_number (n_circ k) p"
       
  2696             using  valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc
       
  2697             apply (intro winding_number_join)
       
  2698             by auto
       
  2699           ultimately show ?thesis using Suc(2) unfolding n_circ_def
       
  2700             by auto
       
  2701         qed
       
  2702       show "pathstart (n_circ (Suc k)) = p + h p"
       
  2703         by (simp add: n_circ_def p_circ_def)
       
  2704       show "pathfinish (n_circ (Suc k)) = p + h p"
       
  2705         using Suc(4) unfolding n_circ_def by auto
       
  2706       show "winding_number (n_circ (Suc k)) p'=0 \<and>  p'\<notin>path_image (n_circ (Suc k))" when "p'\<notin>s-pts" for p'
       
  2707         proof -
       
  2708           have " p' \<notin> path_image p_circ" using \<open>p \<in> s\<close> h p_circ_def that using pcirc_image by blast
       
  2709           moreover have "p' \<notin> path_image (n_circ k)"
       
  2710             using Suc.hyps(7) that by blast
       
  2711           moreover have "winding_number p_circ p' = 0"
       
  2712             proof -
       
  2713               have "path_image p_circ \<subseteq> cball p (h p)"
       
  2714                 using h unfolding p_circ_def using \<open>p \<in> s\<close> by fastforce
       
  2715               moreover have "p'\<notin>cball p (h p)" using \<open>p \<in> s\<close> h that "2.hyps"(2) by fastforce
       
  2716               ultimately show ?thesis unfolding p_circ_def
       
  2717                 apply (intro winding_number_zero_outside)
       
  2718                 by auto
       
  2719             qed
       
  2720           ultimately show ?thesis
       
  2721             unfolding n_Suc
       
  2722             apply (subst winding_number_join)
       
  2723             by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that])
       
  2724         qed
       
  2725       show "f contour_integrable_on (n_circ (Suc k))"
       
  2726         unfolding n_Suc
       
  2727         by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)])
       
  2728       show "contour_integral (n_circ (Suc k)) f = (Suc k) *  contour_integral p_circ f"
       
  2729         unfolding n_Suc
       
  2730         by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]
       
  2731           Suc(9) algebra_simps)
       
  2732     qed
       
  2733   have cp[simp]:"pathstart cp = p + h p"  "pathfinish cp = p + h p"
       
  2734          "valid_path cp" "path_image cp \<subseteq> s - insert p pts"
       
  2735          "winding_number cp p = - n"
       
  2736          "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number cp p'=0 \<and> p' \<notin> path_image cp"
       
  2737          "f contour_integrable_on cp"
       
  2738          "contour_integral cp f = - n * contour_integral p_circ f"
       
  2739     proof -
       
  2740       show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp"
       
  2741         using n_circ unfolding cp_def by auto
       
  2742     next
       
  2743       have "sphere p (h p) \<subseteq>  s - insert p pts"
       
  2744         using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> by force
       
  2745       moreover  have "p + complex_of_real (h p) \<in> s - insert p pts"
       
  2746         using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE)
       
  2747       ultimately show "path_image cp \<subseteq>  s - insert p pts" unfolding cp_def
       
  2748         using n_circ(5)  by auto
       
  2749     next
       
  2750       show "winding_number cp p = - n"
       
  2751         unfolding cp_def using winding_number_reversepath n_circ \<open>h p>0\<close>
       
  2752         by (auto simp: valid_path_imp_path)
       
  2753     next
       
  2754       show "winding_number cp p'=0 \<and> p' \<notin> path_image cp" when "p'\<notin>s - pts" for p'
       
  2755         unfolding cp_def
       
  2756         apply (auto)
       
  2757         apply (subst winding_number_reversepath)
       
  2758         by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1))
       
  2759     next
       
  2760       show "f contour_integrable_on cp" unfolding cp_def
       
  2761         using contour_integrable_reversepath_eq n_circ(1,8) by auto
       
  2762     next
       
  2763       show "contour_integral cp f = - n * contour_integral p_circ f"
       
  2764         unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9)
       
  2765         by auto
       
  2766     qed
       
  2767   define g' where "g' \<equiv> g +++ pg +++ cp +++ (reversepath pg)"
       
  2768   have "contour_integral g' f = (\<Sum>p\<in>pts. winding_number g' p * contour_integral (circlepath p (h p)) f)"
       
  2769     proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \<open>finite pts\<close> ])
       
  2770       show "connected (s - {p} - pts)" using connected by (metis Diff_insert2)
       
  2771       show "open (s - {p})" using \<open>open s\<close> by auto
       
  2772       show " pts \<subseteq> s - {p}" using \<open>insert p pts \<subseteq> s\<close> \<open> p \<notin> pts\<close>  by blast
       
  2773       show "f holomorphic_on s - {p} - pts" using holo \<open>p \<notin> pts\<close> by (metis Diff_insert2)
       
  2774       show "valid_path g'"
       
  2775         unfolding g'_def cp_def using n_circ valid pg g_loop
       
  2776         by (auto intro!:valid_path_join )
       
  2777       show "pathfinish g' = pathstart g'"
       
  2778         unfolding g'_def cp_def using pg(2) by simp
       
  2779       show "path_image g' \<subseteq> s - {p} - pts"
       
  2780         proof -
       
  2781           define s' where "s' \<equiv> s - {p} - pts"
       
  2782           have s':"s' = s-insert p pts " unfolding s'_def by auto
       
  2783           then show ?thesis using path_img pg(4) cp(4)
       
  2784             unfolding g'_def
       
  2785             apply (fold s'_def s')
       
  2786             apply (intro subset_path_image_join)
       
  2787             by auto
       
  2788         qed
       
  2789       note path_join_imp[simp]
       
  2790       show "\<forall>z. z \<notin> s - {p} \<longrightarrow> winding_number g' z = 0"
       
  2791         proof clarify
       
  2792           fix z assume z:"z\<notin>s - {p}"
       
  2793           have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z
       
  2794               + winding_number (pg +++ cp +++ (reversepath pg)) z"
       
  2795             proof (rule winding_number_join)
       
  2796               show "path g" using \<open>valid_path g\<close> by (simp add: valid_path_imp_path)
       
  2797               show "z \<notin> path_image g" using z path_img by auto
       
  2798               show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp
       
  2799                 by (simp add: valid_path_imp_path)
       
  2800             next
       
  2801               have "path_image (pg +++ cp +++ reversepath pg) \<subseteq> s - insert p pts"
       
  2802                 using pg(4) cp(4) by (auto simp:subset_path_image_join)
       
  2803               then show "z \<notin> path_image (pg +++ cp +++ reversepath pg)" using z by auto
       
  2804             next
       
  2805               show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto
       
  2806             qed
       
  2807           also have "... = winding_number g z + (winding_number pg z
       
  2808               + winding_number (cp +++ (reversepath pg)) z)"
       
  2809             proof (subst add_left_cancel,rule winding_number_join)
       
  2810               show "path pg" and "path (cp +++ reversepath pg)"
       
  2811                and "pathfinish pg = pathstart (cp +++ reversepath pg)"
       
  2812                 by (auto simp add: valid_path_imp_path)
       
  2813               show "z \<notin> path_image pg" using pg(4) z by blast
       
  2814               show "z \<notin> path_image (cp +++ reversepath pg)" using z
       
  2815                 by (metis Diff_iff \<open>z \<notin> path_image pg\<close> contra_subsetD cp(4) insertI1
       
  2816                   not_in_path_image_join path_image_reversepath singletonD)
       
  2817             qed
       
  2818           also have "... = winding_number g z + (winding_number pg z
       
  2819               + (winding_number cp z + winding_number (reversepath pg) z))"
       
  2820             apply (auto intro!:winding_number_join simp: valid_path_imp_path)
       
  2821             apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z)
       
  2822             by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z)
       
  2823           also have "... = winding_number g z + winding_number cp z"
       
  2824             apply (subst winding_number_reversepath)
       
  2825             apply (auto simp: valid_path_imp_path)
       
  2826             by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z)
       
  2827           finally have "winding_number g' z = winding_number g z + winding_number cp z"
       
  2828             unfolding g'_def .
       
  2829           moreover have "winding_number g z + winding_number cp z = 0"
       
  2830             using winding z \<open>n=winding_number g p\<close> by auto
       
  2831           ultimately show "winding_number g' z = 0" unfolding g'_def by auto
       
  2832         qed
       
  2833       show "\<forall>pa\<in>s - {p}. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s - {p} \<and> (w \<noteq> pa \<longrightarrow> w \<notin> pts))"
       
  2834         using h by fastforce
       
  2835     qed
       
  2836   moreover have "contour_integral g' f = contour_integral g f
       
  2837       - winding_number g p * contour_integral p_circ f"
       
  2838     proof -
       
  2839       have "contour_integral g' f =  contour_integral g f
       
  2840         + contour_integral (pg +++ cp +++ reversepath pg) f"
       
  2841         unfolding g'_def
       
  2842         apply (subst contour_integral_join)
       
  2843         by (auto simp add:open_Diff[OF \<open>open s\<close>,OF finite_imp_closed[OF fin]]
       
  2844           intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img]
       
  2845           contour_integrable_reversepath)
       
  2846       also have "... = contour_integral g f + contour_integral pg f
       
  2847           + contour_integral (cp +++ reversepath pg) f"
       
  2848         apply (subst contour_integral_join)
       
  2849         by (auto simp add:contour_integrable_reversepath)
       
  2850       also have "... = contour_integral g f + contour_integral pg f
       
  2851           + contour_integral cp f + contour_integral (reversepath pg) f"
       
  2852         apply (subst contour_integral_join)
       
  2853         by (auto simp add:contour_integrable_reversepath)
       
  2854       also have "... = contour_integral g f + contour_integral cp f"
       
  2855         using contour_integral_reversepath
       
  2856         by (auto simp add:contour_integrable_reversepath)
       
  2857       also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f"
       
  2858         using \<open>n=winding_number g p\<close> by auto
       
  2859       finally show ?thesis .
       
  2860     qed
       
  2861   moreover have "winding_number g' p' = winding_number g p'" when "p'\<in>pts" for p'
       
  2862     proof -
       
  2863       have [simp]: "p' \<notin> path_image g" "p' \<notin> path_image pg" "p'\<notin>path_image cp"
       
  2864         using "2.prems"(8) that
       
  2865         apply blast
       
  2866         apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that)
       
  2867         by (meson DiffD2 cp(4) rev_subsetD subset_insertI that)
       
  2868       have "winding_number g' p' = winding_number g p'
       
  2869           + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def
       
  2870         apply (subst winding_number_join)
       
  2871         apply (simp_all add: valid_path_imp_path)
       
  2872         apply (intro not_in_path_image_join)
       
  2873         by auto
       
  2874       also have "... = winding_number g p' + winding_number pg p'
       
  2875           + winding_number (cp +++ reversepath pg) p'"
       
  2876         apply (subst winding_number_join)
       
  2877         apply (simp_all add: valid_path_imp_path)
       
  2878         apply (intro not_in_path_image_join)
       
  2879         by auto
       
  2880       also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p'
       
  2881           + winding_number (reversepath pg) p'"
       
  2882         apply (subst winding_number_join)
       
  2883         by (simp_all add: valid_path_imp_path)
       
  2884       also have "... = winding_number g p' + winding_number cp p'"
       
  2885         apply (subst winding_number_reversepath)
       
  2886         by (simp_all add: valid_path_imp_path)
       
  2887       also have "... = winding_number g p'" using that by auto
       
  2888       finally show ?thesis .
       
  2889     qed
       
  2890   ultimately show ?case unfolding p_circ_def
       
  2891     apply (subst (asm) sum.cong[OF refl,
       
  2892         of pts _ "\<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"])
       
  2893     by (auto simp add:sum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps)
       
  2894 qed
       
  2895 
       
  2896 lemma Cauchy_theorem_singularities:
       
  2897   assumes "open s" "connected s" "finite pts" and
       
  2898           holo:"f holomorphic_on s-pts" and
       
  2899           "valid_path g" and
       
  2900           loop:"pathfinish g = pathstart g" and
       
  2901           "path_image g \<subseteq> s-pts" and
       
  2902           homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0" and
       
  2903           avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
       
  2904   shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
       
  2905     (is "?L=?R")
       
  2906 proof -
       
  2907   define circ where "circ \<equiv> \<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"
       
  2908   define pts1 where "pts1 \<equiv> pts \<inter> s"
       
  2909   define pts2 where "pts2 \<equiv> pts - pts1"
       
  2910   have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> s={}" "pts1\<subseteq>s"
       
  2911     unfolding pts1_def pts2_def by auto
       
  2912   have "contour_integral g f =  (\<Sum>p\<in>pts1. circ p)" unfolding circ_def
       
  2913     proof (rule Cauchy_theorem_aux[OF \<open>open s\<close> _ _ \<open>pts1\<subseteq>s\<close> _ \<open>valid_path g\<close> loop _ homo])
       
  2914       have "finite pts1" unfolding pts1_def using \<open>finite pts\<close> by auto
       
  2915       then show "connected (s - pts1)"
       
  2916         using \<open>open s\<close> \<open>connected s\<close> connected_open_delete_finite[of s] by auto
       
  2917     next
       
  2918       show "finite pts1" using \<open>pts = pts1 \<union> pts2\<close> assms(3) by auto
       
  2919       show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def)
       
  2920       show "path_image g \<subseteq> s - pts1" using assms(7) pts1_def by auto
       
  2921       show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts1))"
       
  2922         by (simp add: avoid pts1_def)
       
  2923     qed
       
  2924   moreover have "sum circ pts2=0"
       
  2925     proof -
       
  2926       have "winding_number g p=0" when "p\<in>pts2" for p
       
  2927         using  \<open>pts2 \<inter> s={}\<close> that homo[rule_format,of p] by auto
       
  2928       thus ?thesis unfolding circ_def
       
  2929         apply (intro sum.neutral)
       
  2930         by auto
       
  2931     qed
       
  2932   moreover have "?R=sum circ pts1 + sum circ pts2"
       
  2933     unfolding circ_def
       
  2934     using sum.union_disjoint[OF _ _ \<open>pts1 \<inter> pts2 = {}\<close>] \<open>finite pts\<close> \<open>pts=pts1 \<union> pts2\<close>
       
  2935     by blast
       
  2936   ultimately show ?thesis
       
  2937     apply (fold circ_def)
       
  2938     by auto
       
  2939 qed
       
  2940 
       
  2941 theorem Residue_theorem:
       
  2942   fixes s pts::"complex set" and f::"complex \<Rightarrow> complex"
       
  2943     and g::"real \<Rightarrow> complex"
       
  2944   assumes "open s" "connected s" "finite pts" and
       
  2945           holo:"f holomorphic_on s-pts" and
       
  2946           "valid_path g" and
       
  2947           loop:"pathfinish g = pathstart g" and
       
  2948           "path_image g \<subseteq> s-pts" and
       
  2949           homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0"
       
  2950   shows "contour_integral g f = 2 * pi * \<i> *(\<Sum>p\<in>pts. winding_number g p * residue f p)"
       
  2951 proof -
       
  2952   define c where "c \<equiv>  2 * pi * \<i>"
       
  2953   obtain h where avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
       
  2954     using finite_cball_avoid[OF \<open>open s\<close> \<open>finite pts\<close>] by metis
       
  2955   have "contour_integral g f
       
  2956       = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
       
  2957     using Cauchy_theorem_singularities[OF assms avoid] .
       
  2958   also have "... = (\<Sum>p\<in>pts.  c * winding_number g p * residue f p)"
       
  2959     proof (intro sum.cong)
       
  2960       show "pts = pts" by simp
       
  2961     next
       
  2962       fix x assume "x \<in> pts"
       
  2963       show "winding_number g x * contour_integral (circlepath x (h x)) f
       
  2964           = c * winding_number g x * residue f x"
       
  2965         proof (cases "x\<in>s")
       
  2966           case False
       
  2967           then have "winding_number g x=0" using homo by auto
       
  2968           thus ?thesis by auto
       
  2969         next
       
  2970           case True
       
  2971           have "contour_integral (circlepath x (h x)) f = c* residue f x"
       
  2972             using \<open>x\<in>pts\<close> \<open>finite pts\<close> avoid[rule_format,OF True]
       
  2973             apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def])
       
  2974             by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \<open>open s\<close> finite_imp_closed])
       
  2975           then show ?thesis by auto
       
  2976         qed
       
  2977     qed
       
  2978   also have "... = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)"
       
  2979     by (simp add: sum_distrib_left algebra_simps)
       
  2980   finally show ?thesis unfolding c_def .
       
  2981 qed
       
  2982 
       
  2983 subsection \<open>Non-essential singular points\<close>
       
  2984 
       
  2985 definition\<^marker>\<open>tag important\<close> is_pole ::
       
  2986   "('a::topological_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool" where
       
  2987   "is_pole f a =  (LIM x (at a). f x :> at_infinity)"
       
  2988 
       
  2989 lemma is_pole_cong:
       
  2990   assumes "eventually (\<lambda>x. f x = g x) (at a)" "a=b"
       
  2991   shows "is_pole f a \<longleftrightarrow> is_pole g b"
       
  2992   unfolding is_pole_def using assms by (intro filterlim_cong,auto)
       
  2993 
       
  2994 lemma is_pole_transform:
       
  2995   assumes "is_pole f a" "eventually (\<lambda>x. f x = g x) (at a)" "a=b"
       
  2996   shows "is_pole g b"
       
  2997   using is_pole_cong assms by auto
       
  2998 
       
  2999 lemma is_pole_tendsto:
       
  3000   fixes f::"('a::topological_space \<Rightarrow> 'b::real_normed_div_algebra)"
       
  3001   shows "is_pole f x \<Longrightarrow> ((inverse o f) \<longlongrightarrow> 0) (at x)"
       
  3002 unfolding is_pole_def
       
  3003 by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at)
       
  3004 
       
  3005 lemma is_pole_inverse_holomorphic:
       
  3006   assumes "open s"
       
  3007     and f_holo:"f holomorphic_on (s-{z})"
       
  3008     and pole:"is_pole f z"
       
  3009     and non_z:"\<forall>x\<in>s-{z}. f x\<noteq>0"
       
  3010   shows "(\<lambda>x. if x=z then 0 else inverse (f x)) holomorphic_on s"
       
  3011 proof -
       
  3012   define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)"
       
  3013   have "isCont g z" unfolding isCont_def  using is_pole_tendsto[OF pole]
       
  3014     apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \<circ> f"])
       
  3015     by (simp_all add:g_def)
       
  3016   moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto
       
  3017   hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def
       
  3018     by (auto elim!:continuous_on_inverse simp add:non_z)
       
  3019   hence "continuous_on (s-{z}) g" unfolding g_def
       
  3020     apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"])
       
  3021     by auto
       
  3022   ultimately have "continuous_on s g" using open_delete[OF \<open>open s\<close>] \<open>open s\<close>
       
  3023     by (auto simp add:continuous_on_eq_continuous_at)
       
  3024   moreover have "(inverse o f) holomorphic_on (s-{z})"
       
  3025     unfolding comp_def using f_holo
       
  3026     by (auto elim!:holomorphic_on_inverse simp add:non_z)
       
  3027   hence "g holomorphic_on (s-{z})"
       
  3028     apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"])
       
  3029     by (auto simp add:g_def)
       
  3030   ultimately show ?thesis unfolding g_def using \<open>open s\<close>
       
  3031     by (auto elim!: no_isolated_singularity)
       
  3032 qed
       
  3033 
       
  3034 lemma not_is_pole_holomorphic:
       
  3035   assumes "open A" "x \<in> A" "f holomorphic_on A"
       
  3036   shows   "\<not>is_pole f x"
       
  3037 proof -
       
  3038   have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact
       
  3039   with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at)
       
  3040   hence "f \<midarrow>x\<rightarrow> f x" by (simp add: isCont_def)
       
  3041   thus "\<not>is_pole f x" unfolding is_pole_def
       
  3042     using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto
       
  3043 qed
       
  3044 
       
  3045 lemma is_pole_inverse_power: "n > 0 \<Longrightarrow> is_pole (\<lambda>z::complex. 1 / (z - a) ^ n) a"
       
  3046   unfolding is_pole_def inverse_eq_divide [symmetric]
       
  3047   by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros)
       
  3048      (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros)
       
  3049 
       
  3050 lemma is_pole_inverse: "is_pole (\<lambda>z::complex. 1 / (z - a)) a"
       
  3051   using is_pole_inverse_power[of 1 a] by simp
       
  3052 
       
  3053 lemma is_pole_divide:
       
  3054   fixes f :: "'a :: t2_space \<Rightarrow> 'b :: real_normed_field"
       
  3055   assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \<noteq> 0"
       
  3056   shows   "is_pole (\<lambda>z. f z / g z) z"
       
  3057 proof -
       
  3058   have "filterlim (\<lambda>z. f z * inverse (g z)) at_infinity (at z)"
       
  3059     by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"]
       
  3060                  filterlim_compose[OF filterlim_inverse_at_infinity])+
       
  3061        (insert assms, auto simp: isCont_def)
       
  3062   thus ?thesis by (simp add: field_split_simps is_pole_def)
       
  3063 qed
       
  3064 
       
  3065 lemma is_pole_basic:
       
  3066   assumes "f holomorphic_on A" "open A" "z \<in> A" "f z \<noteq> 0" "n > 0"
       
  3067   shows   "is_pole (\<lambda>w. f w / (w - z) ^ n) z"
       
  3068 proof (rule is_pole_divide)
       
  3069   have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact
       
  3070   with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at)
       
  3071   have "filterlim (\<lambda>w. (w - z) ^ n) (nhds 0) (at z)"
       
  3072     using assms by (auto intro!: tendsto_eq_intros)
       
  3073   thus "filterlim (\<lambda>w. (w - z) ^ n) (at 0) (at z)"
       
  3074     by (intro filterlim_atI tendsto_eq_intros)
       
  3075        (insert assms, auto simp: eventually_at_filter)
       
  3076 qed fact+
       
  3077 
       
  3078 lemma is_pole_basic':
       
  3079   assumes "f holomorphic_on A" "open A" "0 \<in> A" "f 0 \<noteq> 0" "n > 0"
       
  3080   shows   "is_pole (\<lambda>w. f w / w ^ n) 0"
       
  3081   using is_pole_basic[of f A 0] assms by simp
       
  3082 
       
  3083 text \<open>The proposition
       
  3084               \<^term>\<open>\<exists>x. ((f::complex\<Rightarrow>complex) \<longlongrightarrow> x) (at z) \<or> is_pole f z\<close>
       
  3085 can be interpreted as the complex function \<^term>\<open>f\<close> has a non-essential singularity at \<^term>\<open>z\<close>
       
  3086 (i.e. the singularity is either removable or a pole).\<close>
       
  3087 definition not_essential::"[complex \<Rightarrow> complex, complex] \<Rightarrow> bool" where
       
  3088   "not_essential f z = (\<exists>x. f\<midarrow>z\<rightarrow>x \<or> is_pole f z)"
       
  3089 
       
  3090 definition isolated_singularity_at::"[complex \<Rightarrow> complex, complex] \<Rightarrow> bool" where
       
  3091   "isolated_singularity_at f z = (\<exists>r>0. f analytic_on ball z r-{z})"
       
  3092 
       
  3093 named_theorems singularity_intros "introduction rules for singularities"
       
  3094 
       
  3095 lemma holomorphic_factor_unique:
       
  3096   fixes f::"complex \<Rightarrow> complex" and z::complex and r::real and m n::int
       
  3097   assumes "r>0" "g z\<noteq>0" "h z\<noteq>0"
       
  3098     and asm:"\<forall>w\<in>ball z r-{z}. f w = g w * (w-z) powr n \<and> g w\<noteq>0 \<and> f w =  h w * (w - z) powr m \<and> h w\<noteq>0"
       
  3099     and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r"
       
  3100   shows "n=m"
       
  3101 proof -
       
  3102   have [simp]:"at z within ball z r \<noteq> bot" using \<open>r>0\<close>
       
  3103       by (auto simp add:at_within_ball_bot_iff)
       
  3104   have False when "n>m"
       
  3105   proof -
       
  3106     have "(h \<longlongrightarrow> 0) (at z within ball z r)"
       
  3107     proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) powr (n - m) * g w"])
       
  3108       have "\<forall>w\<in>ball z r-{z}. h w = (w-z)powr(n-m) * g w"
       
  3109         using \<open>n>m\<close> asm \<open>r>0\<close>
       
  3110         apply (auto simp add:field_simps powr_diff)
       
  3111         by force
       
  3112       then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
       
  3113             \<Longrightarrow> (x' - z) powr (n - m) * g x' = h x'" for x' by auto
       
  3114     next
       
  3115       define F where "F \<equiv> at z within ball z r"
       
  3116       define f' where "f' \<equiv> \<lambda>x. (x - z) powr (n-m)"
       
  3117       have "f' z=0" using \<open>n>m\<close> unfolding f'_def by auto
       
  3118       moreover have "continuous F f'" unfolding f'_def F_def continuous_def
       
  3119         apply (subst Lim_ident_at)
       
  3120         using \<open>n>m\<close> by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros)
       
  3121       ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
       
  3122         by (simp add: continuous_within)
       
  3123       moreover have "(g \<longlongrightarrow> g z) F"
       
  3124         using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \<open>r>0\<close>
       
  3125         unfolding F_def by auto
       
  3126       ultimately show " ((\<lambda>w. f' w * g w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
       
  3127     qed
       
  3128     moreover have "(h \<longlongrightarrow> h z) (at z within ball z r)"
       
  3129       using holomorphic_on_imp_continuous_on[OF h_holo]
       
  3130       by (auto simp add:continuous_on_def \<open>r>0\<close>)
       
  3131     ultimately have "h z=0" by (auto intro!: tendsto_unique)
       
  3132     thus False using \<open>h z\<noteq>0\<close> by auto
       
  3133   qed
       
  3134   moreover have False when "m>n"
       
  3135   proof -
       
  3136     have "(g \<longlongrightarrow> 0) (at z within ball z r)"
       
  3137     proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) powr (m - n) * h w"])
       
  3138       have "\<forall>w\<in>ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \<open>m>n\<close> asm
       
  3139         apply (auto simp add:field_simps powr_diff)
       
  3140         by force
       
  3141       then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
       
  3142             \<Longrightarrow> (x' - z) powr (m - n) * h x' = g x'" for x' by auto
       
  3143     next
       
  3144       define F where "F \<equiv> at z within ball z r"
       
  3145       define f' where "f' \<equiv>\<lambda>x. (x - z) powr (m-n)"
       
  3146       have "f' z=0" using \<open>m>n\<close> unfolding f'_def by auto
       
  3147       moreover have "continuous F f'" unfolding f'_def F_def continuous_def
       
  3148         apply (subst Lim_ident_at)
       
  3149         using \<open>m>n\<close> by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros)
       
  3150       ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
       
  3151         by (simp add: continuous_within)
       
  3152       moreover have "(h \<longlongrightarrow> h z) F"
       
  3153         using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \<open>r>0\<close>
       
  3154         unfolding F_def by auto
       
  3155       ultimately show " ((\<lambda>w. f' w * h w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
       
  3156     qed
       
  3157     moreover have "(g \<longlongrightarrow> g z) (at z within ball z r)"
       
  3158       using holomorphic_on_imp_continuous_on[OF g_holo]
       
  3159       by (auto simp add:continuous_on_def \<open>r>0\<close>)
       
  3160     ultimately have "g z=0" by (auto intro!: tendsto_unique)
       
  3161     thus False using \<open>g z\<noteq>0\<close> by auto
       
  3162   qed
       
  3163   ultimately show "n=m" by fastforce
       
  3164 qed
       
  3165 
       
  3166 lemma holomorphic_factor_puncture:
       
  3167   assumes f_iso:"isolated_singularity_at f z"
       
  3168       and "not_essential f z" \<comment> \<open>\<^term>\<open>f\<close> has either a removable singularity or a pole at \<^term>\<open>z\<close>\<close>
       
  3169       and non_zero:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0" \<comment> \<open>\<^term>\<open>f\<close> will not be constantly zero in a neighbour of \<^term>\<open>z\<close>\<close>
       
  3170   shows "\<exists>!n::int. \<exists>g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
       
  3171           \<and> (\<forall>w\<in>cball z r-{z}. f w = g w * (w-z) powr n \<and> g w\<noteq>0)"
       
  3172 proof -
       
  3173   define P where "P = (\<lambda>f n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
       
  3174           \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n)  \<and> g w\<noteq>0))"
       
  3175   have imp_unique:"\<exists>!n::int. \<exists>g r. P f n g r" when "\<exists>n g r. P f n g r"
       
  3176   proof (rule ex_ex1I[OF that])
       
  3177     fix n1 n2 :: int
       
  3178     assume g1_asm:"\<exists>g1 r1. P f n1 g1 r1" and g2_asm:"\<exists>g2 r2. P f n2 g2 r2"
       
  3179     define fac where "fac \<equiv> \<lambda>n g r. \<forall>w\<in>cball z r-{z}. f w = g w * (w - z) powr (of_int n) \<and> g w \<noteq> 0"
       
  3180     obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\<noteq>0"
       
  3181         and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto
       
  3182     obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\<noteq>0"
       
  3183         and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def by auto
       
  3184     define r where "r \<equiv> min r1 r2"
       
  3185     have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto
       
  3186     moreover have "\<forall>w\<in>ball z r-{z}. f w = g1 w * (w-z) powr n1 \<and> g1 w\<noteq>0
       
  3187         \<and> f w = g2 w * (w - z) powr n2  \<and> g2 w\<noteq>0"
       
  3188       using \<open>fac n1 g1 r1\<close> \<open>fac n2 g2 r2\<close>   unfolding fac_def r_def
       
  3189       by fastforce
       
  3190     ultimately show "n1=n2" using g1_holo g2_holo \<open>g1 z\<noteq>0\<close> \<open>g2 z\<noteq>0\<close>
       
  3191       apply (elim holomorphic_factor_unique)
       
  3192       by (auto simp add:r_def)
       
  3193   qed
       
  3194 
       
  3195   have P_exist:"\<exists> n g r. P h n g r" when
       
  3196       "\<exists>z'. (h \<longlongrightarrow> z') (at z)" "isolated_singularity_at h z"  "\<exists>\<^sub>Fw in (at z). h w\<noteq>0"
       
  3197     for h
       
  3198   proof -
       
  3199     from that(2) obtain r where "r>0" "h analytic_on ball z r - {z}"
       
  3200       unfolding isolated_singularity_at_def by auto
       
  3201     obtain z' where "(h \<longlongrightarrow> z') (at z)" using \<open>\<exists>z'. (h \<longlongrightarrow> z') (at z)\<close> by auto
       
  3202     define h' where "h'=(\<lambda>x. if x=z then z' else h x)"
       
  3203     have "h' holomorphic_on ball z r"
       
  3204       apply (rule no_isolated_singularity'[of "{z}"])
       
  3205       subgoal by (metis LIM_equal Lim_at_imp_Lim_at_within \<open>h \<midarrow>z\<rightarrow> z'\<close> empty_iff h'_def insert_iff)
       
  3206       subgoal using \<open>h analytic_on ball z r - {z}\<close> analytic_imp_holomorphic h'_def holomorphic_transform
       
  3207         by fastforce
       
  3208       by auto
       
  3209     have ?thesis when "z'=0"
       
  3210     proof -
       
  3211       have "h' z=0" using that unfolding h'_def by auto
       
  3212       moreover have "\<not> h' constant_on ball z r"
       
  3213         using \<open>\<exists>\<^sub>Fw in (at z). h w\<noteq>0\<close> unfolding constant_on_def frequently_def eventually_at h'_def
       
  3214         apply simp
       
  3215         by (metis \<open>0 < r\<close> centre_in_ball dist_commute mem_ball that)
       
  3216       moreover note \<open>h' holomorphic_on ball z r\<close>
       
  3217       ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \<subseteq> ball z r" and
       
  3218           g:"g holomorphic_on ball z r1"
       
  3219           "\<And>w. w \<in> ball z r1 \<Longrightarrow> h' w = (w - z) ^ n * g w"
       
  3220           "\<And>w. w \<in> ball z r1 \<Longrightarrow> g w \<noteq> 0"
       
  3221         using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified,
       
  3222                 OF \<open>h' holomorphic_on ball z r\<close> \<open>r>0\<close> \<open>h' z=0\<close> \<open>\<not> h' constant_on ball z r\<close>]
       
  3223         by (auto simp add:dist_commute)
       
  3224       define rr where "rr=r1/2"
       
  3225       have "P h' n g rr"
       
  3226         unfolding P_def rr_def
       
  3227         using \<open>n>0\<close> \<open>r1>0\<close> g by (auto simp add:powr_nat)
       
  3228       then have "P h n g rr"
       
  3229         unfolding h'_def P_def by auto
       
  3230       then show ?thesis unfolding P_def by blast
       
  3231     qed
       
  3232     moreover have ?thesis when "z'\<noteq>0"
       
  3233     proof -
       
  3234       have "h' z\<noteq>0" using that unfolding h'_def by auto
       
  3235       obtain r1 where "r1>0" "cball z r1 \<subseteq> ball z r" "\<forall>x\<in>cball z r1. h' x\<noteq>0"
       
  3236       proof -
       
  3237         have "isCont h' z" "h' z\<noteq>0"
       
  3238           by (auto simp add: Lim_cong_within \<open>h \<midarrow>z\<rightarrow> z'\<close> \<open>z'\<noteq>0\<close> continuous_at h'_def)
       
  3239         then obtain r2 where r2:"r2>0" "\<forall>x\<in>ball z r2. h' x\<noteq>0"
       
  3240           using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto
       
  3241         define r1 where "r1=min r2 r / 2"
       
  3242         have "0 < r1" "cball z r1 \<subseteq> ball z r"
       
  3243           using \<open>r2>0\<close> \<open>r>0\<close> unfolding r1_def by auto
       
  3244         moreover have "\<forall>x\<in>cball z r1. h' x \<noteq> 0"
       
  3245           using r2 unfolding r1_def by simp
       
  3246         ultimately show ?thesis using that by auto
       
  3247       qed
       
  3248       then have "P h' 0 h' r1" using \<open>h' holomorphic_on ball z r\<close> unfolding P_def by auto
       
  3249       then have "P h 0 h' r1" unfolding P_def h'_def by auto
       
  3250       then show ?thesis unfolding P_def by blast
       
  3251     qed
       
  3252     ultimately show ?thesis by auto
       
  3253   qed
       
  3254 
       
  3255   have ?thesis when "\<exists>x. (f \<longlongrightarrow> x) (at z)"
       
  3256     apply (rule_tac imp_unique[unfolded P_def])
       
  3257     using P_exist[OF that(1) f_iso non_zero] unfolding P_def .
       
  3258   moreover have ?thesis when "is_pole f z"
       
  3259   proof (rule imp_unique[unfolded P_def])
       
  3260     obtain e where [simp]:"e>0" and e_holo:"f holomorphic_on ball z e - {z}" and e_nz: "\<forall>x\<in>ball z e-{z}. f x\<noteq>0"
       
  3261     proof -
       
  3262       have "\<forall>\<^sub>F z in at z. f z \<noteq> 0"
       
  3263         using \<open>is_pole f z\<close> filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def
       
  3264         by auto
       
  3265       then obtain e1 where e1:"e1>0" "\<forall>x\<in>ball z e1-{z}. f x\<noteq>0"
       
  3266         using that eventually_at[of "\<lambda>x. f x\<noteq>0" z UNIV,simplified] by (auto simp add:dist_commute)
       
  3267       obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}"
       
  3268         using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto
       
  3269       define e where "e=min e1 e2"
       
  3270       show ?thesis
       
  3271         apply (rule that[of e])
       
  3272         using  e1 e2 unfolding e_def by auto
       
  3273     qed
       
  3274 
       
  3275     define h where "h \<equiv> \<lambda>x. inverse (f x)"
       
  3276 
       
  3277     have "\<exists>n g r. P h n g r"
       
  3278     proof -
       
  3279       have "h \<midarrow>z\<rightarrow> 0"
       
  3280         using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce
       
  3281       moreover have "\<exists>\<^sub>Fw in (at z). h w\<noteq>0"
       
  3282         using non_zero
       
  3283         apply (elim frequently_rev_mp)
       
  3284         unfolding h_def eventually_at by (auto intro:exI[where x=1])
       
  3285       moreover have "isolated_singularity_at h z"
       
  3286         unfolding isolated_singularity_at_def h_def
       
  3287         apply (rule exI[where x=e])
       
  3288         using e_holo e_nz \<open>e>0\<close> by (metis open_ball analytic_on_open
       
  3289             holomorphic_on_inverse open_delete)
       
  3290       ultimately show ?thesis
       
  3291         using P_exist[of h] by auto
       
  3292     qed
       
  3293     then obtain n g r
       
  3294       where "0 < r" and
       
  3295             g_holo:"g holomorphic_on cball z r" and "g z\<noteq>0" and
       
  3296             g_fac:"(\<forall>w\<in>cball z r-{z}. h w = g w * (w - z) powr of_int n  \<and> g w \<noteq> 0)"
       
  3297       unfolding P_def by auto
       
  3298     have "P f (-n) (inverse o g) r"
       
  3299     proof -
       
  3300       have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\<in>cball z r - {z}" for w
       
  3301         using g_fac[rule_format,of w] that unfolding h_def
       
  3302         apply (auto simp add:powr_minus )
       
  3303         by (metis inverse_inverse_eq inverse_mult_distrib)
       
  3304       then show ?thesis
       
  3305         unfolding P_def comp_def
       
  3306         using \<open>r>0\<close> g_holo g_fac \<open>g z\<noteq>0\<close> by (auto intro:holomorphic_intros)
       
  3307     qed
       
  3308     then show "\<exists>x g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z \<noteq> 0
       
  3309                   \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int x  \<and> g w \<noteq> 0)"
       
  3310       unfolding P_def by blast
       
  3311   qed
       
  3312   ultimately show ?thesis using \<open>not_essential f z\<close> unfolding not_essential_def  by presburger
       
  3313 qed
       
  3314 
       
  3315 lemma not_essential_transform:
       
  3316   assumes "not_essential g z"
       
  3317   assumes "\<forall>\<^sub>F w in (at z). g w = f w"
       
  3318   shows "not_essential f z"
       
  3319   using assms unfolding not_essential_def
       
  3320   by (simp add: filterlim_cong is_pole_cong)
       
  3321 
       
  3322 lemma isolated_singularity_at_transform:
       
  3323   assumes "isolated_singularity_at g z"
       
  3324   assumes "\<forall>\<^sub>F w in (at z). g w = f w"
       
  3325   shows "isolated_singularity_at f z"
       
  3326 proof -
       
  3327   obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}"
       
  3328     using assms(1) unfolding isolated_singularity_at_def by auto
       
  3329   obtain r2 where "r2>0" and r2:" \<forall>x. x \<noteq> z \<and> dist x z < r2 \<longrightarrow> g x = f x"
       
  3330     using assms(2) unfolding eventually_at by auto
       
  3331   define r3 where "r3=min r1 r2"
       
  3332   have "r3>0" unfolding r3_def using \<open>r1>0\<close> \<open>r2>0\<close> by auto
       
  3333   moreover have "f analytic_on ball z r3 - {z}"
       
  3334   proof -
       
  3335     have "g holomorphic_on ball z r3 - {z}"
       
  3336       using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto)
       
  3337     then have "f holomorphic_on ball z r3 - {z}"
       
  3338       using r2 unfolding r3_def
       
  3339       by (auto simp add:dist_commute elim!:holomorphic_transform)
       
  3340     then show ?thesis by (subst analytic_on_open,auto)
       
  3341   qed
       
  3342   ultimately show ?thesis unfolding isolated_singularity_at_def by auto
       
  3343 qed
       
  3344 
       
  3345 lemma not_essential_powr[singularity_intros]:
       
  3346   assumes "LIM w (at z). f w :> (at x)"
       
  3347   shows "not_essential (\<lambda>w. (f w) powr (of_int n)) z"
       
  3348 proof -
       
  3349   define fp where "fp=(\<lambda>w. (f w) powr (of_int n))"
       
  3350   have ?thesis when "n>0"
       
  3351   proof -
       
  3352     have "(\<lambda>w.  (f w) ^ (nat n)) \<midarrow>z\<rightarrow> x ^ nat n"
       
  3353       using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
       
  3354     then have "fp \<midarrow>z\<rightarrow> x ^ nat n" unfolding fp_def
       
  3355       apply (elim Lim_transform_within[where d=1],simp)
       
  3356       by (metis less_le powr_0 powr_of_int that zero_less_nat_eq zero_power)
       
  3357     then show ?thesis unfolding not_essential_def fp_def by auto
       
  3358   qed
       
  3359   moreover have ?thesis when "n=0"
       
  3360   proof -
       
  3361     have "fp \<midarrow>z\<rightarrow> 1 "
       
  3362       apply (subst tendsto_cong[where g="\<lambda>_.1"])
       
  3363       using that filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def by auto
       
  3364     then show ?thesis unfolding fp_def not_essential_def by auto
       
  3365   qed
       
  3366   moreover have ?thesis when "n<0"
       
  3367   proof (cases "x=0")
       
  3368     case True
       
  3369     have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity"
       
  3370       apply (subst filterlim_inverse_at_iff[symmetric],simp)
       
  3371       apply (rule filterlim_atI)
       
  3372       subgoal using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
       
  3373       subgoal using filterlim_at_within_not_equal[OF assms,of 0]
       
  3374         by (eventually_elim,insert that,auto)
       
  3375       done
       
  3376     then have "LIM w (at z). fp w :> at_infinity"
       
  3377     proof (elim filterlim_mono_eventually)
       
  3378       show "\<forall>\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x"
       
  3379         using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def
       
  3380         apply eventually_elim
       
  3381         using powr_of_int that by auto
       
  3382     qed auto
       
  3383     then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto
       
  3384   next
       
  3385     case False
       
  3386     let ?xx= "inverse (x ^ (nat (-n)))"
       
  3387     have "(\<lambda>w. inverse ((f w) ^ (nat (-n)))) \<midarrow>z\<rightarrow>?xx"
       
  3388       using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
       
  3389     then have "fp \<midarrow>z\<rightarrow>?xx"
       
  3390       apply (elim Lim_transform_within[where d=1],simp)
       
  3391       unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less
       
  3392           not_le power_eq_0_iff powr_0 powr_of_int that)
       
  3393     then show ?thesis unfolding fp_def not_essential_def by auto
       
  3394   qed
       
  3395   ultimately show ?thesis by linarith
       
  3396 qed
       
  3397 
       
  3398 lemma isolated_singularity_at_powr[singularity_intros]:
       
  3399   assumes "isolated_singularity_at f z" "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
       
  3400   shows "isolated_singularity_at (\<lambda>w. (f w) powr (of_int n)) z"
       
  3401 proof -
       
  3402   obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}"
       
  3403     using assms(1) unfolding isolated_singularity_at_def by auto
       
  3404   then have r1:"f holomorphic_on ball z r1 - {z}"
       
  3405     using analytic_on_open[of "ball z r1-{z}" f] by blast
       
  3406   obtain r2 where "r2>0" and r2:"\<forall>w. w \<noteq> z \<and> dist w z < r2 \<longrightarrow> f w \<noteq> 0"
       
  3407     using assms(2) unfolding eventually_at by auto
       
  3408   define r3 where "r3=min r1 r2"
       
  3409   have "(\<lambda>w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}"
       
  3410     apply (rule holomorphic_on_powr_of_int)
       
  3411     subgoal unfolding r3_def using r1 by auto
       
  3412     subgoal unfolding r3_def using r2 by (auto simp add:dist_commute)
       
  3413     done
       
  3414   moreover have "r3>0" unfolding r3_def using \<open>0 < r1\<close> \<open>0 < r2\<close> by linarith
       
  3415   ultimately show ?thesis unfolding isolated_singularity_at_def
       
  3416     apply (subst (asm) analytic_on_open[symmetric])
       
  3417     by auto
       
  3418 qed
       
  3419 
       
  3420 lemma non_zero_neighbour:
       
  3421   assumes f_iso:"isolated_singularity_at f z"
       
  3422       and f_ness:"not_essential f z"
       
  3423       and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
       
  3424     shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
       
  3425 proof -
       
  3426   obtain fn fp fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
       
  3427           and fr: "fp holomorphic_on cball z fr"
       
  3428                   "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
       
  3429     using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto
       
  3430   have "f w \<noteq> 0" when " w \<noteq> z" "dist w z < fr" for w
       
  3431   proof -
       
  3432     have "f w = fp w * (w - z) powr of_int fn" "fp w \<noteq> 0"
       
  3433       using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute)
       
  3434     moreover have "(w - z) powr of_int fn \<noteq>0"
       
  3435       unfolding powr_eq_0_iff using \<open>w\<noteq>z\<close> by auto
       
  3436     ultimately show ?thesis by auto
       
  3437   qed
       
  3438   then show ?thesis using \<open>fr>0\<close> unfolding eventually_at by auto
       
  3439 qed
       
  3440 
       
  3441 lemma non_zero_neighbour_pole:
       
  3442   assumes "is_pole f z"
       
  3443   shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
       
  3444   using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0]
       
  3445   unfolding is_pole_def by auto
       
  3446 
       
  3447 lemma non_zero_neighbour_alt:
       
  3448   assumes holo: "f holomorphic_on S"
       
  3449       and "open S" "connected S" "z \<in> S"  "\<beta> \<in> S" "f \<beta> \<noteq> 0"
       
  3450     shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0 \<and> w\<in>S"
       
  3451 proof (cases "f z = 0")
       
  3452   case True
       
  3453   from isolated_zeros[OF holo \<open>open S\<close> \<open>connected S\<close> \<open>z \<in> S\<close> True \<open>\<beta> \<in> S\<close> \<open>f \<beta> \<noteq> 0\<close>]
       
  3454   obtain r where "0 < r" "ball z r \<subseteq> S" "\<forall>w \<in> ball z r - {z}.f w \<noteq> 0" by metis
       
  3455   then show ?thesis unfolding eventually_at
       
  3456     apply (rule_tac x=r in exI)
       
  3457     by (auto simp add:dist_commute)
       
  3458 next
       
  3459   case False
       
  3460   obtain r1 where r1:"r1>0" "\<forall>y. dist z y < r1 \<longrightarrow> f y \<noteq> 0"
       
  3461     using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at
       
  3462       holo holomorphic_on_imp_continuous_on by blast
       
  3463   obtain r2 where r2:"r2>0" "ball z r2 \<subseteq> S"
       
  3464     using assms(2) assms(4) openE by blast
       
  3465   show ?thesis unfolding eventually_at
       
  3466     apply (rule_tac x="min r1 r2" in exI)
       
  3467     using r1 r2 by (auto simp add:dist_commute)
       
  3468 qed
       
  3469 
       
  3470 lemma not_essential_times[singularity_intros]:
       
  3471   assumes f_ness:"not_essential f z" and g_ness:"not_essential g z"
       
  3472   assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
       
  3473   shows "not_essential (\<lambda>w. f w * g w) z"
       
  3474 proof -
       
  3475   define fg where "fg = (\<lambda>w. f w * g w)"
       
  3476   have ?thesis when "\<not> ((\<exists>\<^sub>Fw in (at z). f w\<noteq>0) \<and> (\<exists>\<^sub>Fw in (at z). g w\<noteq>0))"
       
  3477   proof -
       
  3478     have "\<forall>\<^sub>Fw in (at z). fg w=0"
       
  3479       using that[unfolded frequently_def, simplified] unfolding fg_def
       
  3480       by (auto elim: eventually_rev_mp)
       
  3481     from tendsto_cong[OF this] have "fg \<midarrow>z\<rightarrow>0" by auto
       
  3482     then show ?thesis unfolding not_essential_def fg_def by auto
       
  3483   qed
       
  3484   moreover have ?thesis when f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0" and g_nconst:"\<exists>\<^sub>Fw in (at z). g w\<noteq>0"
       
  3485   proof -
       
  3486     obtain fn fp fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
       
  3487           and fr: "fp holomorphic_on cball z fr"
       
  3488                   "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
       
  3489       using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto
       
  3490     obtain gn gp gr where [simp]:"gp z \<noteq> 0" and "gr > 0"
       
  3491           and gr: "gp holomorphic_on cball z gr"
       
  3492                   "\<forall>w\<in>cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \<and> gp w \<noteq> 0"
       
  3493       using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto
       
  3494 
       
  3495     define r1 where "r1=(min fr gr)"
       
  3496     have "r1>0" unfolding r1_def using  \<open>fr>0\<close> \<open>gr>0\<close> by auto
       
  3497     have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\<noteq>0"
       
  3498       when "w\<in>ball z r1 - {z}" for w
       
  3499     proof -
       
  3500       have "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
       
  3501         using fr(2)[rule_format,of w] that unfolding r1_def by auto
       
  3502       moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \<noteq> 0"
       
  3503         using gr(2)[rule_format, of w] that unfolding r1_def by auto
       
  3504       ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\<noteq>0"
       
  3505         unfolding fg_def by (auto simp add:powr_add)
       
  3506     qed
       
  3507 
       
  3508     have [intro]: "fp \<midarrow>z\<rightarrow>fp z" "gp \<midarrow>z\<rightarrow>gp z"
       
  3509         using fr(1) \<open>fr>0\<close> gr(1) \<open>gr>0\<close>
       
  3510         by (meson open_ball ball_subset_cball centre_in_ball
       
  3511             continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on
       
  3512             holomorphic_on_subset)+
       
  3513     have ?thesis when "fn+gn>0"
       
  3514     proof -
       
  3515       have "(\<lambda>w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \<midarrow>z\<rightarrow>0"
       
  3516         using that by (auto intro!:tendsto_eq_intros)
       
  3517       then have "fg \<midarrow>z\<rightarrow> 0"
       
  3518         apply (elim Lim_transform_within[OF _ \<open>r1>0\<close>])
       
  3519         by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self
       
  3520               eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int
       
  3521               that)
       
  3522       then show ?thesis unfolding not_essential_def fg_def by auto
       
  3523     qed
       
  3524     moreover have ?thesis when "fn+gn=0"
       
  3525     proof -
       
  3526       have "(\<lambda>w. fp w * gp w) \<midarrow>z\<rightarrow>fp z*gp z"
       
  3527         using that by (auto intro!:tendsto_eq_intros)
       
  3528       then have "fg \<midarrow>z\<rightarrow> fp z*gp z"
       
  3529         apply (elim Lim_transform_within[OF _ \<open>r1>0\<close>])
       
  3530         apply (subst fg_times)
       
  3531         by (auto simp add:dist_commute that)
       
  3532       then show ?thesis unfolding not_essential_def fg_def by auto
       
  3533     qed
       
  3534     moreover have ?thesis when "fn+gn<0"
       
  3535     proof -
       
  3536       have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity"
       
  3537         apply (rule filterlim_divide_at_infinity)
       
  3538         apply (insert that, auto intro!:tendsto_eq_intros filterlim_atI)
       
  3539         using eventually_at_topological by blast
       
  3540       then have "is_pole fg z" unfolding is_pole_def
       
  3541         apply (elim filterlim_transform_within[OF _ _ \<open>r1>0\<close>],simp)
       
  3542         apply (subst fg_times,simp add:dist_commute)
       
  3543         apply (subst powr_of_int)
       
  3544         using that by (auto simp add:field_split_simps)
       
  3545       then show ?thesis unfolding not_essential_def fg_def by auto
       
  3546     qed
       
  3547     ultimately show ?thesis unfolding not_essential_def fg_def by fastforce
       
  3548   qed
       
  3549   ultimately show ?thesis by auto
       
  3550 qed
       
  3551 
       
  3552 lemma not_essential_inverse[singularity_intros]:
       
  3553   assumes f_ness:"not_essential f z"
       
  3554   assumes f_iso:"isolated_singularity_at f z"
       
  3555   shows "not_essential (\<lambda>w. inverse (f w)) z"
       
  3556 proof -
       
  3557   define vf where "vf = (\<lambda>w. inverse (f w))"
       
  3558   have ?thesis when "\<not>(\<exists>\<^sub>Fw in (at z). f w\<noteq>0)"
       
  3559   proof -
       
  3560     have "\<forall>\<^sub>Fw in (at z). f w=0"
       
  3561       using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp)
       
  3562     then have "\<forall>\<^sub>Fw in (at z). vf w=0"
       
  3563       unfolding vf_def by auto
       
  3564     from tendsto_cong[OF this] have "vf \<midarrow>z\<rightarrow>0" unfolding vf_def by auto
       
  3565     then show ?thesis unfolding not_essential_def vf_def by auto
       
  3566   qed
       
  3567   moreover have ?thesis when "is_pole f z"
       
  3568   proof -
       
  3569     have "vf \<midarrow>z\<rightarrow>0"
       
  3570       using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast
       
  3571     then show ?thesis unfolding not_essential_def vf_def by auto
       
  3572   qed
       
  3573   moreover have ?thesis when "\<exists>x. f\<midarrow>z\<rightarrow>x " and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
       
  3574   proof -
       
  3575     from that obtain fz where fz:"f\<midarrow>z\<rightarrow>fz" by auto
       
  3576     have ?thesis when "fz=0"
       
  3577     proof -
       
  3578       have "(\<lambda>w. inverse (vf w)) \<midarrow>z\<rightarrow>0"
       
  3579         using fz that unfolding vf_def by auto
       
  3580       moreover have "\<forall>\<^sub>F w in at z. inverse (vf w) \<noteq> 0"
       
  3581         using non_zero_neighbour[OF f_iso f_ness f_nconst]
       
  3582         unfolding vf_def by auto
       
  3583       ultimately have "is_pole vf z"
       
  3584         using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto
       
  3585       then show ?thesis unfolding not_essential_def vf_def by auto
       
  3586     qed
       
  3587     moreover have ?thesis when "fz\<noteq>0"
       
  3588     proof -
       
  3589       have "vf \<midarrow>z\<rightarrow>inverse fz"
       
  3590         using fz that unfolding vf_def by (auto intro:tendsto_eq_intros)
       
  3591       then show ?thesis unfolding not_essential_def vf_def by auto
       
  3592     qed
       
  3593     ultimately show ?thesis by auto
       
  3594   qed
       
  3595   ultimately show ?thesis using f_ness unfolding not_essential_def by auto
       
  3596 qed
       
  3597 
       
  3598 lemma isolated_singularity_at_inverse[singularity_intros]:
       
  3599   assumes f_iso:"isolated_singularity_at f z"
       
  3600       and f_ness:"not_essential f z"
       
  3601   shows "isolated_singularity_at (\<lambda>w. inverse (f w)) z"
       
  3602 proof -
       
  3603   define vf where "vf = (\<lambda>w. inverse (f w))"
       
  3604   have ?thesis when "\<not>(\<exists>\<^sub>Fw in (at z). f w\<noteq>0)"
       
  3605   proof -
       
  3606     have "\<forall>\<^sub>Fw in (at z). f w=0"
       
  3607       using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp)
       
  3608     then have "\<forall>\<^sub>Fw in (at z). vf w=0"
       
  3609       unfolding vf_def by auto
       
  3610     then obtain d1 where "d1>0" and d1:"\<forall>x. x \<noteq> z \<and> dist x z < d1 \<longrightarrow> vf x = 0"
       
  3611       unfolding eventually_at by auto
       
  3612     then have "vf holomorphic_on ball z d1-{z}"
       
  3613       apply (rule_tac holomorphic_transform[of "\<lambda>_. 0"])
       
  3614       by (auto simp add:dist_commute)
       
  3615     then have "vf analytic_on ball z d1 - {z}"
       
  3616       by (simp add: analytic_on_open open_delete)
       
  3617     then show ?thesis using \<open>d1>0\<close> unfolding isolated_singularity_at_def vf_def by auto
       
  3618   qed
       
  3619   moreover have ?thesis when f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
       
  3620   proof -
       
  3621     have "\<forall>\<^sub>F w in at z. f w \<noteq> 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] .
       
  3622     then obtain d1 where d1:"d1>0" "\<forall>x. x \<noteq> z \<and> dist x z < d1 \<longrightarrow> f x \<noteq> 0"
       
  3623       unfolding eventually_at by auto
       
  3624     obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}"
       
  3625       using f_iso unfolding isolated_singularity_at_def by auto
       
  3626     define d3 where "d3=min d1 d2"
       
  3627     have "d3>0" unfolding d3_def using \<open>d1>0\<close> \<open>d2>0\<close> by auto
       
  3628     moreover have "vf analytic_on ball z d3 - {z}"
       
  3629       unfolding vf_def
       
  3630       apply (rule analytic_on_inverse)
       
  3631       subgoal using d2 unfolding d3_def by (elim analytic_on_subset) auto
       
  3632       subgoal for w using d1 unfolding d3_def by (auto simp add:dist_commute)
       
  3633       done
       
  3634     ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto
       
  3635   qed
       
  3636   ultimately show ?thesis by auto
       
  3637 qed
       
  3638 
       
  3639 lemma not_essential_divide[singularity_intros]:
       
  3640   assumes f_ness:"not_essential f z" and g_ness:"not_essential g z"
       
  3641   assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
       
  3642   shows "not_essential (\<lambda>w. f w / g w) z"
       
  3643 proof -
       
  3644   have "not_essential (\<lambda>w. f w * inverse (g w)) z"
       
  3645     apply (rule not_essential_times[where g="\<lambda>w. inverse (g w)"])
       
  3646     using assms by (auto intro: isolated_singularity_at_inverse not_essential_inverse)
       
  3647   then show ?thesis by (simp add:field_simps)
       
  3648 qed
       
  3649 
       
  3650 lemma
       
  3651   assumes f_iso:"isolated_singularity_at f z"
       
  3652       and g_iso:"isolated_singularity_at g z"
       
  3653     shows isolated_singularity_at_times[singularity_intros]:
       
  3654               "isolated_singularity_at (\<lambda>w. f w * g w) z" and
       
  3655           isolated_singularity_at_add[singularity_intros]:
       
  3656               "isolated_singularity_at (\<lambda>w. f w + g w) z"
       
  3657 proof -
       
  3658   obtain d1 d2 where "d1>0" "d2>0"
       
  3659       and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}"
       
  3660     using f_iso g_iso unfolding isolated_singularity_at_def by auto
       
  3661   define d3 where "d3=min d1 d2"
       
  3662   have "d3>0" unfolding d3_def using \<open>d1>0\<close> \<open>d2>0\<close> by auto
       
  3663 
       
  3664   have "(\<lambda>w. f w * g w) analytic_on ball z d3 - {z}"
       
  3665     apply (rule analytic_on_mult)
       
  3666     using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset)
       
  3667   then show "isolated_singularity_at (\<lambda>w. f w * g w) z"
       
  3668     using \<open>d3>0\<close> unfolding isolated_singularity_at_def by auto
       
  3669   have "(\<lambda>w. f w + g w) analytic_on ball z d3 - {z}"
       
  3670     apply (rule analytic_on_add)
       
  3671     using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset)
       
  3672   then show "isolated_singularity_at (\<lambda>w. f w + g w) z"
       
  3673     using \<open>d3>0\<close> unfolding isolated_singularity_at_def by auto
       
  3674 qed
       
  3675 
       
  3676 lemma isolated_singularity_at_uminus[singularity_intros]:
       
  3677   assumes f_iso:"isolated_singularity_at f z"
       
  3678   shows "isolated_singularity_at (\<lambda>w. - f w) z"
       
  3679   using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast
       
  3680 
       
  3681 lemma isolated_singularity_at_id[singularity_intros]:
       
  3682      "isolated_singularity_at (\<lambda>w. w) z"
       
  3683   unfolding isolated_singularity_at_def by (simp add: gt_ex)
       
  3684 
       
  3685 lemma isolated_singularity_at_minus[singularity_intros]:
       
  3686   assumes f_iso:"isolated_singularity_at f z"
       
  3687       and g_iso:"isolated_singularity_at g z"
       
  3688     shows "isolated_singularity_at (\<lambda>w. f w - g w) z"
       
  3689   using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\<lambda>w. - g w"]
       
  3690         ,OF g_iso] by simp
       
  3691 
       
  3692 lemma isolated_singularity_at_divide[singularity_intros]:
       
  3693   assumes f_iso:"isolated_singularity_at f z"
       
  3694       and g_iso:"isolated_singularity_at g z"
       
  3695       and g_ness:"not_essential g z"
       
  3696     shows "isolated_singularity_at (\<lambda>w. f w / g w) z"
       
  3697   using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso,
       
  3698           of "\<lambda>w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps)
       
  3699 
       
  3700 lemma isolated_singularity_at_const[singularity_intros]:
       
  3701     "isolated_singularity_at (\<lambda>w. c) z"
       
  3702   unfolding isolated_singularity_at_def by (simp add: gt_ex)
       
  3703 
       
  3704 lemma isolated_singularity_at_holomorphic:
       
  3705   assumes "f holomorphic_on s-{z}" "open s" "z\<in>s"
       
  3706   shows "isolated_singularity_at f z"
       
  3707   using assms unfolding isolated_singularity_at_def
       
  3708   by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff)
       
  3709 
       
  3710 subsubsection \<open>The order of non-essential singularities (i.e. removable singularities or poles)\<close>
       
  3711 
       
  3712 
       
  3713 definition\<^marker>\<open>tag important\<close> zorder :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> int" where
       
  3714   "zorder f z = (THE n. (\<exists>h r. r>0 \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
       
  3715                    \<and> (\<forall>w\<in>cball z r - {z}. f w =  h w * (w-z) powr (of_int n)
       
  3716                    \<and> h w \<noteq>0)))"
       
  3717 
       
  3718 definition\<^marker>\<open>tag important\<close> zor_poly
       
  3719     ::"[complex \<Rightarrow> complex, complex] \<Rightarrow> complex \<Rightarrow> complex" where
       
  3720   "zor_poly f z = (SOME h. \<exists>r. r > 0 \<and> h holomorphic_on cball z r \<and> h z \<noteq> 0
       
  3721                    \<and> (\<forall>w\<in>cball z r - {z}. f w =  h w * (w - z) powr (zorder f z)
       
  3722                    \<and> h w \<noteq>0))"
       
  3723 
       
  3724 lemma zorder_exist:
       
  3725   fixes f::"complex \<Rightarrow> complex" and z::complex
       
  3726   defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
       
  3727   assumes f_iso:"isolated_singularity_at f z"
       
  3728       and f_ness:"not_essential f z"
       
  3729       and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
       
  3730   shows "g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> g holomorphic_on cball z r
       
  3731     \<and> (\<forall>w\<in>cball z r - {z}. f w  = g w * (w-z) powr n  \<and> g w \<noteq>0))"
       
  3732 proof -
       
  3733   define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
       
  3734           \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n) \<and> g w\<noteq>0))"
       
  3735   have "\<exists>!n. \<exists>g r. P n g r"
       
  3736     using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto
       
  3737   then have "\<exists>g r. P n g r"
       
  3738     unfolding n_def P_def zorder_def
       
  3739     by (drule_tac theI',argo)
       
  3740   then have "\<exists>r. P n g r"
       
  3741     unfolding P_def zor_poly_def g_def n_def
       
  3742     by (drule_tac someI_ex,argo)
       
  3743   then obtain r1 where "P n g r1" by auto
       
  3744   then show ?thesis unfolding P_def by auto
       
  3745 qed
       
  3746 
       
  3747 lemma
       
  3748   fixes f::"complex \<Rightarrow> complex" and z::complex
       
  3749   assumes f_iso:"isolated_singularity_at f z"
       
  3750       and f_ness:"not_essential f z"
       
  3751       and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
       
  3752     shows zorder_inverse: "zorder (\<lambda>w. inverse (f w)) z = - zorder f z"
       
  3753       and zor_poly_inverse: "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. inverse (f w)) z w
       
  3754                                                 = inverse (zor_poly f z w)"
       
  3755 proof -
       
  3756   define vf where "vf = (\<lambda>w. inverse (f w))"
       
  3757   define fn vfn where
       
  3758     "fn = zorder f z"  and "vfn = zorder vf z"
       
  3759   define fp vfp where
       
  3760     "fp = zor_poly f z" and "vfp = zor_poly vf z"
       
  3761 
       
  3762   obtain fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
       
  3763           and fr: "fp holomorphic_on cball z fr"
       
  3764                   "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
       
  3765     using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def]
       
  3766     by auto
       
  3767   have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))"
       
  3768         and fr_nz: "inverse (fp w)\<noteq>0"
       
  3769     when "w\<in>ball z fr - {z}" for w
       
  3770   proof -
       
  3771     have "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
       
  3772       using fr(2)[rule_format,of w] that by auto
       
  3773     then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\<noteq>0"
       
  3774       unfolding vf_def by (auto simp add:powr_minus)
       
  3775   qed
       
  3776   obtain vfr where [simp]:"vfp z \<noteq> 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr"
       
  3777       "(\<forall>w\<in>cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \<and> vfp w \<noteq> 0)"
       
  3778   proof -
       
  3779     have "isolated_singularity_at vf z"
       
  3780       using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def .
       
  3781     moreover have "not_essential vf z"
       
  3782       using not_essential_inverse[OF f_ness f_iso] unfolding vf_def .
       
  3783     moreover have "\<exists>\<^sub>F w in at z. vf w \<noteq> 0"
       
  3784       using f_nconst unfolding vf_def by (auto elim:frequently_elim1)
       
  3785     ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto
       
  3786   qed
       
  3787 
       
  3788 
       
  3789   define r1 where "r1 = min fr vfr"
       
  3790   have "r1>0" using \<open>fr>0\<close> \<open>vfr>0\<close> unfolding r1_def by simp
       
  3791   show "vfn = - fn"
       
  3792     apply (rule holomorphic_factor_unique[of r1 vfp z "\<lambda>w. inverse (fp w)" vf])
       
  3793     subgoal using \<open>r1>0\<close> by simp
       
  3794     subgoal by simp
       
  3795     subgoal by simp
       
  3796     subgoal
       
  3797     proof (rule ballI)
       
  3798       fix w assume "w \<in> ball z r1 - {z}"
       
  3799       then have "w \<in> ball z fr - {z}" "w \<in> cball z vfr - {z}"  unfolding r1_def by auto
       
  3800       from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)]
       
  3801       show "vf w = vfp w * (w - z) powr of_int vfn \<and> vfp w \<noteq> 0
       
  3802               \<and> vf w = inverse (fp w) * (w - z) powr of_int (- fn) \<and> inverse (fp w) \<noteq> 0" by auto
       
  3803     qed
       
  3804     subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros)
       
  3805     subgoal using fr unfolding r1_def by (auto intro!:holomorphic_intros)
       
  3806     done
       
  3807 
       
  3808   have "vfp w = inverse (fp w)" when "w\<in>ball z r1-{z}" for w
       
  3809   proof -
       
  3810     have "w \<in> ball z fr - {z}" "w \<in> cball z vfr - {z}"  "w\<noteq>z" using that unfolding r1_def by auto
       
  3811     from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \<open>vfn = - fn\<close> \<open>w\<noteq>z\<close>
       
  3812     show ?thesis by auto
       
  3813   qed
       
  3814   then show "\<forall>\<^sub>Fw in (at z). vfp w = inverse (fp w)"
       
  3815     unfolding eventually_at using \<open>r1>0\<close>
       
  3816     apply (rule_tac x=r1 in exI)
       
  3817     by (auto simp add:dist_commute)
       
  3818 qed
       
  3819 
       
  3820 lemma
       
  3821   fixes f g::"complex \<Rightarrow> complex" and z::complex
       
  3822   assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
       
  3823       and f_ness:"not_essential f z" and g_ness:"not_essential g z"
       
  3824       and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
       
  3825   shows zorder_times:"zorder (\<lambda>w. f w * g w) z = zorder f z + zorder g z" and
       
  3826         zor_poly_times:"\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w * g w) z w
       
  3827                                                   = zor_poly f z w *zor_poly g z w"
       
  3828 proof -
       
  3829   define fg where "fg = (\<lambda>w. f w * g w)"
       
  3830   define fn gn fgn where
       
  3831     "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z"
       
  3832   define fp gp fgp where
       
  3833     "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z"
       
  3834   have f_nconst:"\<exists>\<^sub>Fw in (at z). f w \<noteq> 0" and g_nconst:"\<exists>\<^sub>Fw in (at z).g w\<noteq> 0"
       
  3835     using fg_nconst by (auto elim!:frequently_elim1)
       
  3836   obtain fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
       
  3837           and fr: "fp holomorphic_on cball z fr"
       
  3838                   "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
       
  3839     using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto
       
  3840   obtain gr where [simp]:"gp z \<noteq> 0" and "gr > 0"
       
  3841           and gr: "gp holomorphic_on cball z gr"
       
  3842                   "\<forall>w\<in>cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \<and> gp w \<noteq> 0"
       
  3843     using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto
       
  3844   define r1 where "r1=min fr gr"
       
  3845   have "r1>0" unfolding r1_def using \<open>fr>0\<close> \<open>gr>0\<close> by auto
       
  3846   have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\<noteq>0"
       
  3847     when "w\<in>ball z r1 - {z}" for w
       
  3848   proof -
       
  3849     have "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
       
  3850       using fr(2)[rule_format,of w] that unfolding r1_def by auto
       
  3851     moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \<noteq> 0"
       
  3852       using gr(2)[rule_format, of w] that unfolding r1_def by auto
       
  3853     ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\<noteq>0"
       
  3854       unfolding fg_def by (auto simp add:powr_add)
       
  3855   qed
       
  3856 
       
  3857   obtain fgr where [simp]:"fgp z \<noteq> 0" and "fgr > 0"
       
  3858           and fgr: "fgp holomorphic_on cball z fgr"
       
  3859                   "\<forall>w\<in>cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0"
       
  3860   proof -
       
  3861     have "fgp z \<noteq> 0 \<and> (\<exists>r>0. fgp holomorphic_on cball z r
       
  3862             \<and> (\<forall>w\<in>cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0))"
       
  3863       apply (rule zorder_exist[of fg z, folded fgn_def fgp_def])
       
  3864       subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] .
       
  3865       subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] .
       
  3866       subgoal unfolding fg_def using fg_nconst .
       
  3867       done
       
  3868     then show ?thesis using that by blast
       
  3869   qed
       
  3870   define r2 where "r2 = min fgr r1"
       
  3871   have "r2>0" using \<open>r1>0\<close> \<open>fgr>0\<close> unfolding r2_def by simp
       
  3872   show "fgn = fn + gn "
       
  3873     apply (rule holomorphic_factor_unique[of r2 fgp z "\<lambda>w. fp w * gp w" fg])
       
  3874     subgoal using \<open>r2>0\<close> by simp
       
  3875     subgoal by simp
       
  3876     subgoal by simp
       
  3877     subgoal
       
  3878     proof (rule ballI)
       
  3879       fix w assume "w \<in> ball z r2 - {z}"
       
  3880       then have "w \<in> ball z r1 - {z}" "w \<in> cball z fgr - {z}"  unfolding r2_def by auto
       
  3881       from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)]
       
  3882       show "fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0
       
  3883               \<and> fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \<and> fp w * gp w \<noteq> 0" by auto
       
  3884     qed
       
  3885     subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros)
       
  3886     subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros)
       
  3887     done
       
  3888 
       
  3889   have "fgp w = fp w *gp w" when "w\<in>ball z r2-{z}" for w
       
  3890   proof -
       
  3891     have "w \<in> ball z r1 - {z}" "w \<in> cball z fgr - {z}" "w\<noteq>z" using that  unfolding r2_def by auto
       
  3892     from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \<open>fgn = fn + gn\<close> \<open>w\<noteq>z\<close>
       
  3893     show ?thesis by auto
       
  3894   qed
       
  3895   then show "\<forall>\<^sub>Fw in (at z). fgp w = fp w * gp w"
       
  3896     using \<open>r2>0\<close> unfolding eventually_at by (auto simp add:dist_commute)
       
  3897 qed
       
  3898 
       
  3899 lemma
       
  3900   fixes f g::"complex \<Rightarrow> complex" and z::complex
       
  3901   assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
       
  3902       and f_ness:"not_essential f z" and g_ness:"not_essential g z"
       
  3903       and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
       
  3904   shows zorder_divide:"zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z" and
       
  3905         zor_poly_divide:"\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w / g w) z w
       
  3906                                                   = zor_poly f z w  / zor_poly g z w"
       
  3907 proof -
       
  3908   have f_nconst:"\<exists>\<^sub>Fw in (at z). f w \<noteq> 0" and g_nconst:"\<exists>\<^sub>Fw in (at z).g w\<noteq> 0"
       
  3909     using fg_nconst by (auto elim!:frequently_elim1)
       
  3910   define vg where "vg=(\<lambda>w. inverse (g w))"
       
  3911   have "zorder (\<lambda>w. f w * vg w) z = zorder f z + zorder vg z"
       
  3912     apply (rule zorder_times[OF f_iso _ f_ness,of vg])
       
  3913     subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] .
       
  3914     subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] .
       
  3915     subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1)
       
  3916     done
       
  3917   then show "zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z"
       
  3918     using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
       
  3919     by (auto simp add:field_simps)
       
  3920 
       
  3921   have "\<forall>\<^sub>F w in at z. zor_poly (\<lambda>w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w"
       
  3922     apply (rule zor_poly_times[OF f_iso _ f_ness,of vg])
       
  3923     subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] .
       
  3924     subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] .
       
  3925     subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1)
       
  3926     done
       
  3927   then show "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w / g w) z w = zor_poly f z w  / zor_poly g z w"
       
  3928     using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
       
  3929     apply eventually_elim
       
  3930     by (auto simp add:field_simps)
       
  3931 qed
       
  3932 
       
  3933 lemma zorder_exist_zero:
       
  3934   fixes f::"complex \<Rightarrow> complex" and z::complex
       
  3935   defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
       
  3936   assumes  holo: "f holomorphic_on s" and
       
  3937           "open s" "connected s" "z\<in>s"
       
  3938       and non_const: "\<exists>w\<in>s. f w \<noteq> 0"
       
  3939   shows "(if f z=0 then n > 0 else n=0) \<and> (\<exists>r. r>0 \<and> cball z r \<subseteq> s \<and> g holomorphic_on cball z r
       
  3940     \<and> (\<forall>w\<in>cball z r. f w  = g w * (w-z) ^ nat n  \<and> g w \<noteq>0))"
       
  3941 proof -
       
  3942   obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> s" "g holomorphic_on cball z r"
       
  3943             "(\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
       
  3944   proof -
       
  3945     have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r
       
  3946             \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0))"
       
  3947     proof (rule zorder_exist[of f z,folded g_def n_def])
       
  3948       show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
       
  3949         using holo assms(4,6)
       
  3950         by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE)
       
  3951       show "not_essential f z" unfolding not_essential_def
       
  3952         using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on
       
  3953         by fastforce
       
  3954       have "\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w\<in>s"
       
  3955       proof -
       
  3956         obtain w where "w\<in>s" "f w\<noteq>0" using non_const by auto
       
  3957         then show ?thesis
       
  3958           by (rule non_zero_neighbour_alt[OF holo \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close>])
       
  3959       qed
       
  3960       then show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
       
  3961         apply (elim eventually_frequentlyE)
       
  3962         by auto
       
  3963     qed
       
  3964     then obtain r1 where "g z \<noteq> 0" "r1>0" and r1:"g holomorphic_on cball z r1"
       
  3965             "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
       
  3966       by auto
       
  3967     obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> s"
       
  3968       using assms(4,6) open_contains_cball_eq by blast
       
  3969     define r3 where "r3=min r1 r2"
       
  3970     have "r3>0" "cball z r3 \<subseteq> s" using \<open>r1>0\<close> r2 unfolding r3_def by auto
       
  3971     moreover have "g holomorphic_on cball z r3"
       
  3972       using r1(1) unfolding r3_def by auto
       
  3973     moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
       
  3974       using r1(2) unfolding r3_def by auto
       
  3975     ultimately show ?thesis using that[of r3] \<open>g z\<noteq>0\<close> by auto
       
  3976   qed
       
  3977 
       
  3978   have if_0:"if f z=0 then n > 0 else n=0"
       
  3979   proof -
       
  3980     have "f\<midarrow> z \<rightarrow> f z"
       
  3981       by (metis assms(4,6,7) at_within_open continuous_on holo holomorphic_on_imp_continuous_on)
       
  3982     then have "(\<lambda>w. g w * (w - z) powr of_int n) \<midarrow>z\<rightarrow> f z"
       
  3983       apply (elim Lim_transform_within_open[where s="ball z r"])
       
  3984       using r by auto
       
  3985     moreover have "g \<midarrow>z\<rightarrow>g z"
       
  3986       by (metis (mono_tags, lifting) open_ball at_within_open_subset
       
  3987           ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1,3) subsetCE)
       
  3988     ultimately have "(\<lambda>w. (g w * (w - z) powr of_int n) / g w) \<midarrow>z\<rightarrow> f z/g z"
       
  3989       apply (rule_tac tendsto_divide)
       
  3990       using \<open>g z\<noteq>0\<close> by auto
       
  3991     then have powr_tendsto:"(\<lambda>w. (w - z) powr of_int n) \<midarrow>z\<rightarrow> f z/g z"
       
  3992       apply (elim Lim_transform_within_open[where s="ball z r"])
       
  3993       using r by auto
       
  3994 
       
  3995     have ?thesis when "n\<ge>0" "f z=0"
       
  3996     proof -
       
  3997       have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> f z/g z"
       
  3998         using powr_tendsto
       
  3999         apply (elim Lim_transform_within[where d=r])
       
  4000         by (auto simp add: powr_of_int \<open>n\<ge>0\<close> \<open>r>0\<close>)
       
  4001       then have *:"(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 0" using \<open>f z=0\<close> by simp
       
  4002       moreover have False when "n=0"
       
  4003       proof -
       
  4004         have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 1"
       
  4005           using \<open>n=0\<close> by auto
       
  4006         then show False using * using LIM_unique zero_neq_one by blast
       
  4007       qed
       
  4008       ultimately show ?thesis using that by fastforce
       
  4009     qed
       
  4010     moreover have ?thesis when "n\<ge>0" "f z\<noteq>0"
       
  4011     proof -
       
  4012       have False when "n>0"
       
  4013       proof -
       
  4014         have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> f z/g z"
       
  4015           using powr_tendsto
       
  4016           apply (elim Lim_transform_within[where d=r])
       
  4017           by (auto simp add: powr_of_int \<open>n\<ge>0\<close> \<open>r>0\<close>)
       
  4018         moreover have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 0"
       
  4019           using \<open>n>0\<close> by (auto intro!:tendsto_eq_intros)
       
  4020         ultimately show False using \<open>f z\<noteq>0\<close> \<open>g z\<noteq>0\<close> using LIM_unique divide_eq_0_iff by blast
       
  4021       qed
       
  4022       then show ?thesis using that by force
       
  4023     qed
       
  4024     moreover have False when "n<0"
       
  4025     proof -
       
  4026       have "(\<lambda>w. inverse ((w - z) ^ nat (- n))) \<midarrow>z\<rightarrow> f z/g z"
       
  4027            "(\<lambda>w.((w - z) ^ nat (- n))) \<midarrow>z\<rightarrow> 0"
       
  4028         subgoal  using powr_tendsto powr_of_int that
       
  4029           by (elim Lim_transform_within_open[where s=UNIV],auto)
       
  4030         subgoal using that by (auto intro!:tendsto_eq_intros)
       
  4031         done
       
  4032       from tendsto_mult[OF this,simplified]
       
  4033       have "(\<lambda>x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \<midarrow>z\<rightarrow> 0" .
       
  4034       then have "(\<lambda>x. 1::complex) \<midarrow>z\<rightarrow> 0"
       
  4035         by (elim Lim_transform_within_open[where s=UNIV],auto)
       
  4036       then show False using LIM_const_eq by fastforce
       
  4037     qed
       
  4038     ultimately show ?thesis by fastforce
       
  4039   qed
       
  4040   moreover have "f w  = g w * (w-z) ^ nat n  \<and> g w \<noteq>0" when "w\<in>cball z r" for w
       
  4041   proof (cases "w=z")
       
  4042     case True
       
  4043     then have "f \<midarrow>z\<rightarrow>f w"
       
  4044       using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce
       
  4045     then have "(\<lambda>w. g w * (w-z) ^ nat n) \<midarrow>z\<rightarrow>f w"
       
  4046     proof (elim Lim_transform_within[OF _ \<open>r>0\<close>])
       
  4047       fix x assume "0 < dist x z" "dist x z < r"
       
  4048       then have "x \<in> cball z r - {z}" "x\<noteq>z"
       
  4049         unfolding cball_def by (auto simp add: dist_commute)
       
  4050       then have "f x = g x * (x - z) powr of_int n"
       
  4051         using r(4)[rule_format,of x] by simp
       
  4052       also have "... = g x * (x - z) ^ nat n"
       
  4053         apply (subst powr_of_int)
       
  4054         using if_0 \<open>x\<noteq>z\<close> by (auto split:if_splits)
       
  4055       finally show "f x = g x * (x - z) ^ nat n" .
       
  4056     qed
       
  4057     moreover have "(\<lambda>w. g w * (w-z) ^ nat n) \<midarrow>z\<rightarrow> g w * (w-z) ^ nat n"
       
  4058       using True apply (auto intro!:tendsto_eq_intros)
       
  4059       by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball
       
  4060           continuous_on holomorphic_on_imp_continuous_on r(1) r(3) that)
       
  4061     ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast
       
  4062     then show ?thesis using \<open>g z\<noteq>0\<close> True by auto
       
  4063   next
       
  4064     case False
       
  4065     then have "f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0"
       
  4066       using r(4) that by auto
       
  4067     then show ?thesis using False if_0 powr_of_int by (auto split:if_splits)
       
  4068   qed
       
  4069   ultimately show ?thesis using r by auto
       
  4070 qed
       
  4071 
       
  4072 lemma zorder_exist_pole:
       
  4073   fixes f::"complex \<Rightarrow> complex" and z::complex
       
  4074   defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
       
  4075   assumes  holo: "f holomorphic_on s-{z}" and
       
  4076           "open s" "z\<in>s"
       
  4077       and "is_pole f z"
       
  4078   shows "n < 0 \<and> g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> cball z r \<subseteq> s \<and> g holomorphic_on cball z r
       
  4079     \<and> (\<forall>w\<in>cball z r - {z}. f w  = g w / (w-z) ^ nat (- n) \<and> g w \<noteq>0))"
       
  4080 proof -
       
  4081   obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> s" "g holomorphic_on cball z r"
       
  4082             "(\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
       
  4083   proof -
       
  4084     have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r
       
  4085             \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0))"
       
  4086     proof (rule zorder_exist[of f z,folded g_def n_def])
       
  4087       show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
       
  4088         using holo assms(4,5)
       
  4089         by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff)
       
  4090       show "not_essential f z" unfolding not_essential_def
       
  4091         using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on
       
  4092         by fastforce
       
  4093       from non_zero_neighbour_pole[OF \<open>is_pole f z\<close>] show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
       
  4094         apply (elim eventually_frequentlyE)
       
  4095         by auto
       
  4096     qed
       
  4097     then obtain r1 where "g z \<noteq> 0" "r1>0" and r1:"g holomorphic_on cball z r1"
       
  4098             "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
       
  4099       by auto
       
  4100     obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> s"
       
  4101       using assms(4,5) open_contains_cball_eq by metis
       
  4102     define r3 where "r3=min r1 r2"
       
  4103     have "r3>0" "cball z r3 \<subseteq> s" using \<open>r1>0\<close> r2 unfolding r3_def by auto
       
  4104     moreover have "g holomorphic_on cball z r3"
       
  4105       using r1(1) unfolding r3_def by auto
       
  4106     moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
       
  4107       using r1(2) unfolding r3_def by auto
       
  4108     ultimately show ?thesis using that[of r3] \<open>g z\<noteq>0\<close> by auto
       
  4109   qed
       
  4110 
       
  4111   have "n<0"
       
  4112   proof (rule ccontr)
       
  4113     assume " \<not> n < 0"
       
  4114     define c where "c=(if n=0 then g z else 0)"
       
  4115     have [simp]:"g \<midarrow>z\<rightarrow> g z"
       
  4116       by (metis open_ball at_within_open ball_subset_cball centre_in_ball
       
  4117             continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) )
       
  4118     have "\<forall>\<^sub>F x in at z. f x = g x * (x - z) ^ nat n"
       
  4119       unfolding eventually_at_topological
       
  4120       apply (rule_tac exI[where x="ball z r"])
       
  4121       using r powr_of_int \<open>\<not> n < 0\<close> by auto
       
  4122     moreover have "(\<lambda>x. g x * (x - z) ^ nat n) \<midarrow>z\<rightarrow>c"
       
  4123     proof (cases "n=0")
       
  4124       case True
       
  4125       then show ?thesis unfolding c_def by simp
       
  4126     next
       
  4127       case False
       
  4128       then have "(\<lambda>x. (x - z) ^ nat n) \<midarrow>z\<rightarrow> 0" using \<open>\<not> n < 0\<close>
       
  4129         by (auto intro!:tendsto_eq_intros)
       
  4130       from tendsto_mult[OF _ this,of g "g z",simplified]
       
  4131       show ?thesis unfolding c_def using False by simp
       
  4132     qed
       
  4133     ultimately have "f \<midarrow>z\<rightarrow>c" using tendsto_cong by fast
       
  4134     then show False using \<open>is_pole f z\<close> at_neq_bot not_tendsto_and_filterlim_at_infinity
       
  4135       unfolding is_pole_def by blast
       
  4136   qed
       
  4137   moreover have "\<forall>w\<in>cball z r - {z}. f w  = g w / (w-z) ^ nat (- n) \<and> g w \<noteq>0"
       
  4138     using r(4) \<open>n<0\<close> powr_of_int
       
  4139     by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le)
       
  4140   ultimately show ?thesis using r(1-3) \<open>g z\<noteq>0\<close> by auto
       
  4141 qed
       
  4142 
       
  4143 lemma zorder_eqI:
       
  4144   assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0"
       
  4145   assumes fg_eq:"\<And>w. \<lbrakk>w \<in> s;w\<noteq>z\<rbrakk> \<Longrightarrow> f w = g w * (w - z) powr n"
       
  4146   shows   "zorder f z = n"
       
  4147 proof -
       
  4148   have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact
       
  4149   moreover have "open (-{0::complex})" by auto
       
  4150   ultimately have "open ((g -` (-{0})) \<inter> s)"
       
  4151     unfolding continuous_on_open_vimage[OF \<open>open s\<close>] by blast
       
  4152   moreover from assms have "z \<in> (g -` (-{0})) \<inter> s" by auto
       
  4153   ultimately obtain r where r: "r > 0" "cball z r \<subseteq>  s \<inter> (g -` (-{0}))"
       
  4154     unfolding open_contains_cball by blast
       
  4155 
       
  4156   let ?gg= "(\<lambda>w. g w * (w - z) powr n)"
       
  4157   define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
       
  4158           \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n) \<and> g w\<noteq>0))"
       
  4159   have "P n g r"
       
  4160     unfolding P_def using r assms(3,4,5) by auto
       
  4161   then have "\<exists>g r. P n g r" by auto
       
  4162   moreover have unique: "\<exists>!n. \<exists>g r. P n g r" unfolding P_def
       
  4163   proof (rule holomorphic_factor_puncture)
       
  4164     have "ball z r-{z} \<subseteq> s" using r using ball_subset_cball by blast
       
  4165     then have "?gg holomorphic_on ball z r-{z}"
       
  4166       using \<open>g holomorphic_on s\<close> r by (auto intro!: holomorphic_intros)
       
  4167     then have "f holomorphic_on ball z r - {z}"
       
  4168       apply (elim holomorphic_transform)
       
  4169       using fg_eq \<open>ball z r-{z} \<subseteq> s\<close> by auto
       
  4170     then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
       
  4171       using analytic_on_open open_delete r(1) by blast
       
  4172   next
       
  4173     have "not_essential ?gg z"
       
  4174     proof (intro singularity_intros)
       
  4175       show "not_essential g z"
       
  4176         by (meson \<open>continuous_on s g\<close> assms(1) assms(2) continuous_on_eq_continuous_at
       
  4177             isCont_def not_essential_def)
       
  4178       show " \<forall>\<^sub>F w in at z. w - z \<noteq> 0" by (simp add: eventually_at_filter)
       
  4179       then show "LIM w at z. w - z :> at 0"
       
  4180         unfolding filterlim_at by (auto intro:tendsto_eq_intros)
       
  4181       show "isolated_singularity_at g z"
       
  4182         by (meson Diff_subset open_ball analytic_on_holomorphic
       
  4183             assms(1,2,3) holomorphic_on_subset isolated_singularity_at_def openE)
       
  4184     qed
       
  4185     then show "not_essential f z"
       
  4186       apply (elim not_essential_transform)
       
  4187       unfolding eventually_at using assms(1,2) assms(5)[symmetric]
       
  4188       by (metis dist_commute mem_ball openE subsetCE)
       
  4189     show "\<exists>\<^sub>F w in at z. f w \<noteq> 0" unfolding frequently_at
       
  4190     proof (rule,rule)
       
  4191       fix d::real assume "0 < d"
       
  4192       define z' where "z'=z+min d r / 2"
       
  4193       have "z' \<noteq> z" " dist z' z < d "
       
  4194         unfolding z'_def using \<open>d>0\<close> \<open>r>0\<close>
       
  4195         by (auto simp add:dist_norm)
       
  4196       moreover have "f z' \<noteq> 0"
       
  4197       proof (subst fg_eq[OF _ \<open>z'\<noteq>z\<close>])
       
  4198         have "z' \<in> cball z r" unfolding z'_def using \<open>r>0\<close> \<open>d>0\<close> by (auto simp add:dist_norm)
       
  4199         then show " z' \<in> s" using r(2) by blast
       
  4200         show "g z' * (z' - z) powr of_int n \<noteq> 0"
       
  4201           using P_def \<open>P n g r\<close> \<open>z' \<in> cball z r\<close> calculation(1) by auto
       
  4202       qed
       
  4203       ultimately show "\<exists>x\<in>UNIV. x \<noteq> z \<and> dist x z < d \<and> f x \<noteq> 0" by auto
       
  4204     qed
       
  4205   qed
       
  4206   ultimately have "(THE n. \<exists>g r. P n g r) = n"
       
  4207     by (rule_tac the1_equality)
       
  4208   then show ?thesis unfolding zorder_def P_def by blast
       
  4209 qed
       
  4210 
       
  4211 lemma residue_pole_order:
       
  4212   fixes f::"complex \<Rightarrow> complex" and z::complex
       
  4213   defines "n \<equiv> nat (- zorder f z)" and "h \<equiv> zor_poly f z"
       
  4214   assumes f_iso:"isolated_singularity_at f z"
       
  4215     and pole:"is_pole f z"
       
  4216   shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))"
       
  4217 proof -
       
  4218   define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)"
       
  4219   obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}"
       
  4220     using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast
       
  4221   obtain r where "0 < n" "0 < r" and r_cball:"cball z r \<subseteq> ball z e" and h_holo: "h holomorphic_on cball z r"
       
  4222       and h_divide:"(\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w - z) ^ n) \<and> h w \<noteq> 0)"
       
  4223   proof -
       
  4224     obtain r where r:"zorder f z < 0" "h z \<noteq> 0" "r>0" "cball z r \<subseteq> ball z e" "h holomorphic_on cball z r"
       
  4225         "(\<forall>w\<in>cball z r - {z}. f w = h w / (w - z) ^ n \<and> h w \<noteq> 0)"
       
  4226       using zorder_exist_pole[OF f_holo,simplified,OF \<open>is_pole f z\<close>,folded n_def h_def] by auto
       
  4227     have "n>0" using \<open>zorder f z < 0\<close> unfolding n_def by simp
       
  4228     moreover have "(\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w - z) ^ n) \<and> h w \<noteq> 0)"
       
  4229       using \<open>h z\<noteq>0\<close> r(6) by blast
       
  4230     ultimately show ?thesis using r(3,4,5) that by blast
       
  4231   qed
       
  4232   have r_nonzero:"\<And>w. w \<in> ball z r - {z} \<Longrightarrow> f w \<noteq> 0"
       
  4233     using h_divide by simp
       
  4234   define c where "c \<equiv> 2 * pi * \<i>"
       
  4235   define der_f where "der_f \<equiv> ((deriv ^^ (n - 1)) h z / fact (n-1))"
       
  4236   define h' where "h' \<equiv> \<lambda>u. h u / (u - z) ^ n"
       
  4237   have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)"
       
  4238     unfolding h'_def
       
  4239     proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1",
       
  4240         folded c_def Suc_pred'[OF \<open>n>0\<close>]])
       
  4241       show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp
       
  4242       show "h holomorphic_on ball z r" using h_holo by auto
       
  4243       show " z \<in> ball z r" using \<open>r>0\<close> by auto
       
  4244     qed
       
  4245   then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto
       
  4246   then have "(f has_contour_integral c * der_f) (circlepath z r)"
       
  4247     proof (elim has_contour_integral_eq)
       
  4248       fix x assume "x \<in> path_image (circlepath z r)"
       
  4249       hence "x\<in>cball z r - {z}" using \<open>r>0\<close> by auto
       
  4250       then show "h' x = f x" using h_divide unfolding h'_def by auto
       
  4251     qed
       
  4252   moreover have "(f has_contour_integral c * residue f z) (circlepath z r)"
       
  4253     using base_residue[of \<open>ball z e\<close> z,simplified,OF \<open>r>0\<close> f_holo r_cball,folded c_def]
       
  4254     unfolding c_def by simp
       
  4255   ultimately have "c * der_f =  c * residue f z" using has_contour_integral_unique by blast
       
  4256   hence "der_f = residue f z" unfolding c_def by auto
       
  4257   thus ?thesis unfolding der_f_def by auto
       
  4258 qed
       
  4259 
       
  4260 lemma simple_zeroI:
       
  4261   assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0"
       
  4262   assumes "\<And>w. w \<in> s \<Longrightarrow> f w = g w * (w - z)"
       
  4263   shows   "zorder f z = 1"
       
  4264   using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto)
       
  4265 
       
  4266 lemma higher_deriv_power:
       
  4267   shows   "(deriv ^^ j) (\<lambda>w. (w - z) ^ n) w =
       
  4268              pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)"
       
  4269 proof (induction j arbitrary: w)
       
  4270   case 0
       
  4271   thus ?case by auto
       
  4272 next
       
  4273   case (Suc j w)
       
  4274   have "(deriv ^^ Suc j) (\<lambda>w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\<lambda>w. (w - z) ^ n)) w"
       
  4275     by simp
       
  4276   also have "(deriv ^^ j) (\<lambda>w. (w - z) ^ n) =
       
  4277                (\<lambda>w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))"
       
  4278     using Suc by (intro Suc.IH ext)
       
  4279   also {
       
  4280     have "(\<dots> has_field_derivative of_nat (n - j) *
       
  4281                pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)"
       
  4282       using Suc.prems by (auto intro!: derivative_eq_intros)
       
  4283     also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j =
       
  4284                  pochhammer (of_nat (Suc n - Suc j)) (Suc j)"
       
  4285       by (cases "Suc j \<le> n", subst pochhammer_rec)
       
  4286          (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left)
       
  4287     finally have "deriv (\<lambda>w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w =
       
  4288                     \<dots> * (w - z) ^ (n - Suc j)"
       
  4289       by (rule DERIV_imp_deriv)
       
  4290   }
       
  4291   finally show ?case .
       
  4292 qed
       
  4293 
       
  4294 lemma zorder_zero_eqI:
       
  4295   assumes  f_holo:"f holomorphic_on s" and "open s" "z \<in> s"
       
  4296   assumes zero: "\<And>i. i < nat n \<Longrightarrow> (deriv ^^ i) f z = 0"
       
  4297   assumes nz: "(deriv ^^ nat n) f z \<noteq> 0" and "n\<ge>0"
       
  4298   shows   "zorder f z = n"
       
  4299 proof -
       
  4300   obtain r where [simp]:"r>0" and "ball z r \<subseteq> s"
       
  4301     using \<open>open s\<close> \<open>z\<in>s\<close> openE by blast
       
  4302   have nz':"\<exists>w\<in>ball z r. f w \<noteq> 0"
       
  4303   proof (rule ccontr)
       
  4304     assume "\<not> (\<exists>w\<in>ball z r. f w \<noteq> 0)"
       
  4305     then have "eventually (\<lambda>u. f u = 0) (nhds z)"
       
  4306       using \<open>r>0\<close> unfolding eventually_nhds
       
  4307       apply (rule_tac x="ball z r" in exI)
       
  4308       by auto
       
  4309     then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\<lambda>_. 0) z"
       
  4310       by (intro higher_deriv_cong_ev) auto
       
  4311     also have "(deriv ^^ nat n) (\<lambda>_. 0) z = 0"
       
  4312       by (induction n) simp_all
       
  4313     finally show False using nz by contradiction
       
  4314   qed
       
  4315 
       
  4316   define zn g where "zn = zorder f z" and "g = zor_poly f z"
       
  4317   obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and
       
  4318             [simp]:"e>0" and "cball z e \<subseteq> ball z r" and
       
  4319             g_holo:"g holomorphic_on cball z e" and
       
  4320             e_fac:"(\<forall>w\<in>cball z e. f w = g w * (w - z) ^ nat zn \<and> g w \<noteq> 0)"
       
  4321   proof -
       
  4322     have "f holomorphic_on ball z r"
       
  4323       using f_holo \<open>ball z r \<subseteq> s\<close> by auto
       
  4324     from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def]
       
  4325     show ?thesis by blast
       
  4326   qed
       
  4327   from this(1,2,5) have "zn\<ge>0" "g z\<noteq>0"
       
  4328     subgoal by (auto split:if_splits)
       
  4329     subgoal using \<open>0 < e\<close> ball_subset_cball centre_in_ball e_fac by blast
       
  4330     done
       
  4331 
       
  4332   define A where "A = (\<lambda>i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)"
       
  4333   have deriv_A:"(deriv ^^ i) f z = (if zn \<le> int i then A i else 0)" for i
       
  4334   proof -
       
  4335     have "eventually (\<lambda>w. w \<in> ball z e) (nhds z)"
       
  4336       using \<open>cball z e \<subseteq> ball z r\<close> \<open>e>0\<close> by (intro eventually_nhds_in_open) auto
       
  4337     hence "eventually (\<lambda>w. f w = (w - z) ^ (nat zn) * g w) (nhds z)"
       
  4338       apply eventually_elim
       
  4339       by (use e_fac in auto)
       
  4340     hence "(deriv ^^ i) f z = (deriv ^^ i) (\<lambda>w. (w - z) ^ nat zn * g w) z"
       
  4341       by (intro higher_deriv_cong_ev) auto
       
  4342     also have "\<dots> = (\<Sum>j=0..i. of_nat (i choose j) *
       
  4343                        (deriv ^^ j) (\<lambda>w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)"
       
  4344       using g_holo \<open>e>0\<close>
       
  4345       by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros)
       
  4346     also have "\<dots> = (\<Sum>j=0..i. if j = nat zn then
       
  4347                     of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)"
       
  4348     proof (intro sum.cong refl, goal_cases)
       
  4349       case (1 j)
       
  4350       have "(deriv ^^ j) (\<lambda>w. (w - z) ^ nat zn) z =
       
  4351               pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)"
       
  4352         by (subst higher_deriv_power) auto
       
  4353       also have "\<dots> = (if j = nat zn then fact j else 0)"
       
  4354         by (auto simp: not_less pochhammer_0_left pochhammer_fact)
       
  4355       also have "of_nat (i choose j) * \<dots> * (deriv ^^ (i - j)) g z =
       
  4356                    (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn)
       
  4357                         * (deriv ^^ (i - nat zn)) g z else 0)"
       
  4358         by simp
       
  4359       finally show ?case .
       
  4360     qed
       
  4361     also have "\<dots> = (if i \<ge> zn then A i else 0)"
       
  4362       by (auto simp: A_def)
       
  4363     finally show "(deriv ^^ i) f z = \<dots>" .
       
  4364   qed
       
  4365 
       
  4366   have False when "n<zn"
       
  4367   proof -
       
  4368     have "(deriv ^^ nat n) f z = 0"
       
  4369       using deriv_A[of "nat n"] that \<open>n\<ge>0\<close> by auto
       
  4370     with nz show False by auto
       
  4371   qed
       
  4372   moreover have "n\<le>zn"
       
  4373   proof -
       
  4374     have "g z \<noteq> 0" using e_fac[rule_format,of z] \<open>e>0\<close> by simp
       
  4375     then have "(deriv ^^ nat zn) f z \<noteq> 0"
       
  4376       using deriv_A[of "nat zn"] by(auto simp add:A_def)
       
  4377     then have "nat zn \<ge> nat n" using zero[of "nat zn"] by linarith
       
  4378     moreover have "zn\<ge>0" using e_if by (auto split:if_splits)
       
  4379     ultimately show ?thesis using nat_le_eq_zle by blast
       
  4380   qed
       
  4381   ultimately show ?thesis unfolding zn_def by fastforce
       
  4382 qed
       
  4383 
       
  4384 lemma
       
  4385   assumes "eventually (\<lambda>z. f z = g z) (at z)" "z = z'"
       
  4386   shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'"
       
  4387 proof -
       
  4388   define P where "P = (\<lambda>ff n h r. 0 < r \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
       
  4389                     \<and> (\<forall>w\<in>cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \<and> h w\<noteq>0))"
       
  4390   have "(\<exists>r. P f n h r) = (\<exists>r. P g n h r)" for n h
       
  4391   proof -
       
  4392     have *: "\<exists>r. P g n h r" if "\<exists>r. P f n h r" and "eventually (\<lambda>x. f x = g x) (at z)" for f g
       
  4393     proof -
       
  4394       from that(1) obtain r1 where r1_P:"P f n h r1" by auto
       
  4395       from that(2) obtain r2 where "r2>0" and r2_dist:"\<forall>x. x \<noteq> z \<and> dist x z \<le> r2 \<longrightarrow> f x = g x"
       
  4396         unfolding eventually_at_le by auto
       
  4397       define r where "r=min r1 r2"
       
  4398       have "r>0" "h z\<noteq>0" using r1_P \<open>r2>0\<close> unfolding r_def P_def by auto
       
  4399       moreover have "h holomorphic_on cball z r"
       
  4400         using r1_P unfolding P_def r_def by auto
       
  4401       moreover have "g w = h w * (w - z) powr of_int n \<and> h w \<noteq> 0" when "w\<in>cball z r - {z}" for w
       
  4402       proof -
       
  4403         have "f w = h w * (w - z) powr of_int n \<and> h w \<noteq> 0"
       
  4404           using r1_P that unfolding P_def r_def by auto
       
  4405         moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def
       
  4406           by (simp add: dist_commute)
       
  4407         ultimately show ?thesis by simp
       
  4408       qed
       
  4409       ultimately show ?thesis unfolding P_def by auto
       
  4410     qed
       
  4411     from assms have eq': "eventually (\<lambda>z. g z = f z) (at z)"
       
  4412       by (simp add: eq_commute)
       
  4413     show ?thesis
       
  4414       by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']])
       
  4415   qed
       
  4416   then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'"
       
  4417       using \<open>z=z'\<close> unfolding P_def zorder_def zor_poly_def by auto
       
  4418 qed
       
  4419 
       
  4420 lemma zorder_nonzero_div_power:
       
  4421   assumes "open s" "z \<in> s" "f holomorphic_on s" "f z \<noteq> 0" "n > 0"
       
  4422   shows  "zorder (\<lambda>w. f w / (w - z) ^ n) z = - n"
       
  4423   apply (rule zorder_eqI[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>f holomorphic_on s\<close> \<open>f z\<noteq>0\<close>])
       
  4424   apply (subst powr_of_int)
       
  4425   using \<open>n>0\<close> by (auto simp add:field_simps)
       
  4426 
       
  4427 lemma zor_poly_eq:
       
  4428   assumes "isolated_singularity_at f z" "not_essential f z" "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
       
  4429   shows "eventually (\<lambda>w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)"
       
  4430 proof -
       
  4431   obtain r where r:"r>0"
       
  4432        "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))"
       
  4433     using zorder_exist[OF assms] by blast
       
  4434   then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z"
       
  4435     by (auto simp: field_simps powr_minus)
       
  4436   have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
       
  4437     using r eventually_at_ball'[of r z UNIV] by auto
       
  4438   thus ?thesis by eventually_elim (insert *, auto)
       
  4439 qed
       
  4440 
       
  4441 lemma zor_poly_zero_eq:
       
  4442   assumes "f holomorphic_on s" "open s" "connected s" "z \<in> s" "\<exists>w\<in>s. f w \<noteq> 0"
       
  4443   shows "eventually (\<lambda>w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)"
       
  4444 proof -
       
  4445   obtain r where r:"r>0"
       
  4446        "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))"
       
  4447     using zorder_exist_zero[OF assms] by auto
       
  4448   then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)"
       
  4449     by (auto simp: field_simps powr_minus)
       
  4450   have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
       
  4451     using r eventually_at_ball'[of r z UNIV] by auto
       
  4452   thus ?thesis by eventually_elim (insert *, auto)
       
  4453 qed
       
  4454 
       
  4455 lemma zor_poly_pole_eq:
       
  4456   assumes f_iso:"isolated_singularity_at f z" "is_pole f z"
       
  4457   shows "eventually (\<lambda>w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)"
       
  4458 proof -
       
  4459   obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}"
       
  4460     using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast
       
  4461   obtain r where r:"r>0"
       
  4462        "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))"
       
  4463     using zorder_exist_pole[OF f_holo,simplified,OF \<open>is_pole f z\<close>] by auto
       
  4464   then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)"
       
  4465     by (auto simp: field_simps)
       
  4466   have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
       
  4467     using r eventually_at_ball'[of r z UNIV] by auto
       
  4468   thus ?thesis by eventually_elim (insert *, auto)
       
  4469 qed
       
  4470 
       
  4471 lemma zor_poly_eqI:
       
  4472   fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
       
  4473   defines "n \<equiv> zorder f z0"
       
  4474   assumes "isolated_singularity_at f z0" "not_essential f z0" "\<exists>\<^sub>F w in at z0. f w \<noteq> 0"
       
  4475   assumes lim: "((\<lambda>x. f (g x) * (g x - z0) powr - n) \<longlongrightarrow> c) F"
       
  4476   assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
       
  4477   shows   "zor_poly f z0 z0 = c"
       
  4478 proof -
       
  4479   from zorder_exist[OF assms(2-4)] obtain r where
       
  4480     r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r"
       
  4481        "\<And>w. w \<in> cball z0 r - {z0} \<Longrightarrow> f w = zor_poly f z0 w * (w - z0) powr n"
       
  4482     unfolding n_def by blast
       
  4483   from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
       
  4484     using eventually_at_ball'[of r z0 UNIV] by auto
       
  4485   hence "eventually (\<lambda>w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)"
       
  4486     by eventually_elim (insert r, auto simp: field_simps powr_minus)
       
  4487   moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
       
  4488     using r by (intro holomorphic_on_imp_continuous_on) auto
       
  4489   with r(1,2) have "isCont (zor_poly f z0) z0"
       
  4490     by (auto simp: continuous_on_eq_continuous_at)
       
  4491   hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
       
  4492     unfolding isCont_def .
       
  4493   ultimately have "((\<lambda>w. f w * (w - z0) powr - n) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
       
  4494     by (blast intro: Lim_transform_eventually)
       
  4495   hence "((\<lambda>x. f (g x) * (g x - z0) powr - n) \<longlongrightarrow> zor_poly f z0 z0) F"
       
  4496     by (rule filterlim_compose[OF _ g])
       
  4497   from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
       
  4498 qed
       
  4499 
       
  4500 lemma zor_poly_zero_eqI:
       
  4501   fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
       
  4502   defines "n \<equiv> zorder f z0"
       
  4503   assumes "f holomorphic_on A" "open A" "connected A" "z0 \<in> A" "\<exists>z\<in>A. f z \<noteq> 0"
       
  4504   assumes lim: "((\<lambda>x. f (g x) / (g x - z0) ^ nat n) \<longlongrightarrow> c) F"
       
  4505   assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
       
  4506   shows   "zor_poly f z0 z0 = c"
       
  4507 proof -
       
  4508   from zorder_exist_zero[OF assms(2-6)] obtain r where
       
  4509     r: "r > 0" "cball z0 r \<subseteq> A" "zor_poly f z0 holomorphic_on cball z0 r"
       
  4510        "\<And>w. w \<in> cball z0 r \<Longrightarrow> f w = zor_poly f z0 w * (w - z0) ^ nat n"
       
  4511     unfolding n_def by blast
       
  4512   from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
       
  4513     using eventually_at_ball'[of r z0 UNIV] by auto
       
  4514   hence "eventually (\<lambda>w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)"
       
  4515     by eventually_elim (insert r, auto simp: field_simps)
       
  4516   moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
       
  4517     using r by (intro holomorphic_on_imp_continuous_on) auto
       
  4518   with r(1,2) have "isCont (zor_poly f z0) z0"
       
  4519     by (auto simp: continuous_on_eq_continuous_at)
       
  4520   hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
       
  4521     unfolding isCont_def .
       
  4522   ultimately have "((\<lambda>w. f w / (w - z0) ^ nat n) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
       
  4523     by (blast intro: Lim_transform_eventually)
       
  4524   hence "((\<lambda>x. f (g x) / (g x - z0) ^ nat n) \<longlongrightarrow> zor_poly f z0 z0) F"
       
  4525     by (rule filterlim_compose[OF _ g])
       
  4526   from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
       
  4527 qed
       
  4528 
       
  4529 lemma zor_poly_pole_eqI:
       
  4530   fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
       
  4531   defines "n \<equiv> zorder f z0"
       
  4532   assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0"
       
  4533   assumes lim: "((\<lambda>x. f (g x) * (g x - z0) ^ nat (-n)) \<longlongrightarrow> c) F"
       
  4534   assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
       
  4535   shows   "zor_poly f z0 z0 = c"
       
  4536 proof -
       
  4537   obtain r where r: "r > 0"  "zor_poly f z0 holomorphic_on cball z0 r"
       
  4538   proof -
       
  4539     have "\<exists>\<^sub>F w in at z0. f w \<noteq> 0"
       
  4540       using non_zero_neighbour_pole[OF \<open>is_pole f z0\<close>] by (auto elim:eventually_frequentlyE)
       
  4541     moreover have "not_essential f z0" unfolding not_essential_def using \<open>is_pole f z0\<close> by simp
       
  4542     ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto
       
  4543   qed
       
  4544   from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
       
  4545     using eventually_at_ball'[of r z0 UNIV] by auto
       
  4546   have "eventually (\<lambda>w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)"
       
  4547     using zor_poly_pole_eq[OF f_iso \<open>is_pole f z0\<close>] unfolding n_def .
       
  4548   moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
       
  4549     using r by (intro holomorphic_on_imp_continuous_on) auto
       
  4550   with r(1,2) have "isCont (zor_poly f z0) z0"
       
  4551     by (auto simp: continuous_on_eq_continuous_at)
       
  4552   hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
       
  4553     unfolding isCont_def .
       
  4554   ultimately have "((\<lambda>w. f w * (w - z0) ^ nat (-n)) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
       
  4555     by (blast intro: Lim_transform_eventually)
       
  4556   hence "((\<lambda>x. f (g x) * (g x - z0) ^ nat (-n)) \<longlongrightarrow> zor_poly f z0 z0) F"
       
  4557     by (rule filterlim_compose[OF _ g])
       
  4558   from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
       
  4559 qed
       
  4560 
       
  4561 lemma residue_simple_pole:
       
  4562   assumes "isolated_singularity_at f z0"
       
  4563   assumes "is_pole f z0" "zorder f z0 = - 1"
       
  4564   shows   "residue f z0 = zor_poly f z0 z0"
       
  4565   using assms by (subst residue_pole_order) simp_all
       
  4566 
       
  4567 lemma residue_simple_pole_limit:
       
  4568   assumes "isolated_singularity_at f z0"
       
  4569   assumes "is_pole f z0" "zorder f z0 = - 1"
       
  4570   assumes "((\<lambda>x. f (g x) * (g x - z0)) \<longlongrightarrow> c) F"
       
  4571   assumes "filterlim g (at z0) F" "F \<noteq> bot"
       
  4572   shows   "residue f z0 = c"
       
  4573 proof -
       
  4574   have "residue f z0 = zor_poly f z0 z0"
       
  4575     by (rule residue_simple_pole assms)+
       
  4576   also have "\<dots> = c"
       
  4577     apply (rule zor_poly_pole_eqI)
       
  4578     using assms by auto
       
  4579   finally show ?thesis .
       
  4580 qed
       
  4581 
       
  4582 lemma lhopital_complex_simple:
       
  4583   assumes "(f has_field_derivative f') (at z)"
       
  4584   assumes "(g has_field_derivative g') (at z)"
       
  4585   assumes "f z = 0" "g z = 0" "g' \<noteq> 0" "f' / g' = c"
       
  4586   shows   "((\<lambda>w. f w / g w) \<longlongrightarrow> c) (at z)"
       
  4587 proof -
       
  4588   have "eventually (\<lambda>w. w \<noteq> z) (at z)"
       
  4589     by (auto simp: eventually_at_filter)
       
  4590   hence "eventually (\<lambda>w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)"
       
  4591     by eventually_elim (simp add: assms field_split_simps)
       
  4592   moreover have "((\<lambda>w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \<longlongrightarrow> f' / g') (at z)"
       
  4593     by (intro tendsto_divide has_field_derivativeD assms)
       
  4594   ultimately have "((\<lambda>w. f w / g w) \<longlongrightarrow> f' / g') (at z)"
       
  4595     by (blast intro: Lim_transform_eventually)
       
  4596   with assms show ?thesis by simp
       
  4597 qed
       
  4598 
       
  4599 lemma
       
  4600   assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s"
       
  4601           and "open s" "connected s" "z \<in> s"
       
  4602   assumes g_deriv:"(g has_field_derivative g') (at z)"
       
  4603   assumes "f z \<noteq> 0" "g z = 0" "g' \<noteq> 0"
       
  4604   shows   porder_simple_pole_deriv: "zorder (\<lambda>w. f w / g w) z = - 1"
       
  4605     and   residue_simple_pole_deriv: "residue (\<lambda>w. f w / g w) z = f z / g'"
       
  4606 proof -
       
  4607   have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z"
       
  4608     using isolated_singularity_at_holomorphic[OF _ \<open>open s\<close> \<open>z\<in>s\<close>] f_holo g_holo
       
  4609     by (meson Diff_subset holomorphic_on_subset)+
       
  4610   have [simp]:"not_essential f z" "not_essential g z"
       
  4611     unfolding not_essential_def using f_holo g_holo assms(3,5)
       
  4612     by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+
       
  4613   have g_nconst:"\<exists>\<^sub>F w in at z. g w \<noteq>0 "
       
  4614   proof (rule ccontr)
       
  4615     assume "\<not> (\<exists>\<^sub>F w in at z. g w \<noteq> 0)"
       
  4616     then have "\<forall>\<^sub>F w in nhds z. g w = 0"
       
  4617       unfolding eventually_at eventually_nhds frequently_at using \<open>g z = 0\<close>
       
  4618       by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball)
       
  4619     then have "deriv g z = deriv (\<lambda>_. 0) z"
       
  4620       by (intro deriv_cong_ev) auto
       
  4621     then have "deriv g z = 0" by auto
       
  4622     then have "g' = 0" using g_deriv DERIV_imp_deriv by blast
       
  4623     then show False using \<open>g'\<noteq>0\<close> by auto
       
  4624   qed
       
  4625 
       
  4626   have "zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z"
       
  4627   proof -
       
  4628     have "\<forall>\<^sub>F w in at z. f w \<noteq>0 \<and> w\<in>s"
       
  4629       apply (rule non_zero_neighbour_alt)
       
  4630       using assms by auto
       
  4631     with g_nconst have "\<exists>\<^sub>F w in at z. f w * g w \<noteq> 0"
       
  4632       by (elim frequently_rev_mp eventually_rev_mp,auto)
       
  4633     then show ?thesis using zorder_divide[of f z g] by auto
       
  4634   qed
       
  4635   moreover have "zorder f z=0"
       
  4636     apply (rule zorder_zero_eqI[OF f_holo \<open>open s\<close> \<open>z\<in>s\<close>])
       
  4637     using \<open>f z\<noteq>0\<close> by auto
       
  4638   moreover have "zorder g z=1"
       
  4639     apply (rule zorder_zero_eqI[OF g_holo \<open>open s\<close> \<open>z\<in>s\<close>])
       
  4640     subgoal using assms(8) by auto
       
  4641     subgoal using DERIV_imp_deriv assms(9) g_deriv by auto
       
  4642     subgoal by simp
       
  4643     done
       
  4644   ultimately show "zorder (\<lambda>w. f w / g w) z = - 1" by auto
       
  4645 
       
  4646   show "residue (\<lambda>w. f w / g w) z = f z / g'"
       
  4647   proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified])
       
  4648     show "zorder (\<lambda>w. f w / g w) z = - 1" by fact
       
  4649     show "isolated_singularity_at (\<lambda>w. f w / g w) z"
       
  4650       by (auto intro: singularity_intros)
       
  4651     show "is_pole (\<lambda>w. f w / g w) z"
       
  4652     proof (rule is_pole_divide)
       
  4653       have "\<forall>\<^sub>F x in at z. g x \<noteq> 0"
       
  4654         apply (rule non_zero_neighbour)
       
  4655         using g_nconst by auto
       
  4656       moreover have "g \<midarrow>z\<rightarrow> 0"
       
  4657         using DERIV_isCont assms(8) continuous_at g_deriv by force
       
  4658       ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp
       
  4659       show "isCont f z"
       
  4660         using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on
       
  4661         by auto
       
  4662       show "f z \<noteq> 0" by fact
       
  4663     qed
       
  4664     show "filterlim id (at z) (at z)" by (simp add: filterlim_iff)
       
  4665     have "((\<lambda>w. (f w * (w - z)) / g w) \<longlongrightarrow> f z / g') (at z)"
       
  4666     proof (rule lhopital_complex_simple)
       
  4667       show "((\<lambda>w. f w * (w - z)) has_field_derivative f z) (at z)"
       
  4668         using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo])
       
  4669       show "(g has_field_derivative g') (at z)" by fact
       
  4670     qed (insert assms, auto)
       
  4671     then show "((\<lambda>w. (f w / g w) * (w - z)) \<longlongrightarrow> f z / g') (at z)"
       
  4672       by (simp add: field_split_simps)
       
  4673   qed
       
  4674 qed
       
  4675 
       
  4676 subsection \<open>The argument principle\<close>
       
  4677 
       
  4678 theorem argument_principle:
       
  4679   fixes f::"complex \<Rightarrow> complex" and poles s:: "complex set"
       
  4680   defines "pz \<equiv> {w. f w = 0 \<or> w \<in> poles}" \<comment> \<open>\<^term>\<open>pz\<close> is the set of poles and zeros\<close>
       
  4681   assumes "open s" and
       
  4682           "connected s" and
       
  4683           f_holo:"f holomorphic_on s-poles" and
       
  4684           h_holo:"h holomorphic_on s" and
       
  4685           "valid_path g" and
       
  4686           loop:"pathfinish g = pathstart g" and
       
  4687           path_img:"path_image g \<subseteq> s - pz" and
       
  4688           homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and
       
  4689           finite:"finite pz" and
       
  4690           poles:"\<forall>p\<in>poles. is_pole f p"
       
  4691   shows "contour_integral g (\<lambda>x. deriv f x * h x / f x) = 2 * pi * \<i> *
       
  4692           (\<Sum>p\<in>pz. winding_number g p * h p * zorder f p)"
       
  4693     (is "?L=?R")
       
  4694 proof -
       
  4695   define c where "c \<equiv> 2 * complex_of_real pi * \<i> "
       
  4696   define ff where "ff \<equiv> (\<lambda>x. deriv f x * h x / f x)"
       
  4697   define cont where "cont \<equiv> \<lambda>ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)"
       
  4698   define avoid where "avoid \<equiv> \<lambda>p e. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pz)"
       
  4699 
       
  4700   have "\<exists>e>0. avoid p e \<and> (p\<in>pz \<longrightarrow> cont ff p e)" when "p\<in>s" for p
       
  4701   proof -
       
  4702     obtain e1 where "e1>0" and e1_avoid:"avoid p e1"
       
  4703       using finite_cball_avoid[OF \<open>open s\<close> finite] \<open>p\<in>s\<close> unfolding avoid_def by auto
       
  4704     have "\<exists>e2>0. cball p e2 \<subseteq> ball p e1 \<and> cont ff p e2" when "p\<in>pz"
       
  4705     proof -
       
  4706       define po where "po \<equiv> zorder f p"
       
  4707       define pp where "pp \<equiv> zor_poly f p"
       
  4708       define f' where "f' \<equiv> \<lambda>w. pp w * (w - p) powr po"
       
  4709       define ff' where "ff' \<equiv> (\<lambda>x. deriv f' x * h x / f' x)"
       
  4710       obtain r where "pp p\<noteq>0" "r>0" and
       
  4711           "r<e1" and
       
  4712           pp_holo:"pp holomorphic_on cball p r" and
       
  4713           pp_po:"(\<forall>w\<in>cball p r-{p}. f w = pp w * (w - p) powr po \<and> pp w \<noteq> 0)"
       
  4714       proof -
       
  4715         have "isolated_singularity_at f p"
       
  4716         proof -
       
  4717           have "f holomorphic_on ball p e1 - {p}"
       
  4718             apply (intro holomorphic_on_subset[OF f_holo])
       
  4719             using e1_avoid \<open>p\<in>pz\<close> unfolding avoid_def pz_def by force
       
  4720           then show ?thesis unfolding isolated_singularity_at_def
       
  4721             using \<open>e1>0\<close> analytic_on_open open_delete by blast
       
  4722         qed
       
  4723         moreover have "not_essential f p"
       
  4724         proof (cases "is_pole f p")
       
  4725           case True
       
  4726           then show ?thesis unfolding not_essential_def by auto
       
  4727         next
       
  4728           case False
       
  4729           then have "p\<in>s-poles" using \<open>p\<in>s\<close> poles unfolding pz_def by auto
       
  4730           moreover have "open (s-poles)"
       
  4731             using \<open>open s\<close>
       
  4732             apply (elim open_Diff)
       
  4733             apply (rule finite_imp_closed)
       
  4734             using finite unfolding pz_def by simp
       
  4735           ultimately have "isCont f p"
       
  4736             using holomorphic_on_imp_continuous_on[OF f_holo] continuous_on_eq_continuous_at
       
  4737             by auto
       
  4738           then show ?thesis unfolding isCont_def not_essential_def by auto
       
  4739         qed
       
  4740         moreover have "\<exists>\<^sub>F w in at p. f w \<noteq> 0 "
       
  4741         proof (rule ccontr)
       
  4742           assume "\<not> (\<exists>\<^sub>F w in at p. f w \<noteq> 0)"
       
  4743           then have "\<forall>\<^sub>F w in at p. f w= 0" unfolding frequently_def by auto
       
  4744           then obtain rr where "rr>0" "\<forall>w\<in>ball p rr - {p}. f w =0"
       
  4745             unfolding eventually_at by (auto simp add:dist_commute)
       
  4746           then have "ball p rr - {p} \<subseteq> {w\<in>ball p rr-{p}. f w=0}" by blast
       
  4747           moreover have "infinite (ball p rr - {p})" using \<open>rr>0\<close> using finite_imp_not_open by fastforce
       
  4748           ultimately have "infinite {w\<in>ball p rr-{p}. f w=0}" using infinite_super by blast
       
  4749           then have "infinite pz"
       
  4750             unfolding pz_def infinite_super by auto
       
  4751           then show False using \<open>finite pz\<close> by auto
       
  4752         qed
       
  4753         ultimately obtain r where "pp p \<noteq> 0" and r:"r>0" "pp holomorphic_on cball p r"
       
  4754                   "(\<forall>w\<in>cball p r - {p}. f w = pp w * (w - p) powr of_int po \<and> pp w \<noteq> 0)"
       
  4755           using zorder_exist[of f p,folded po_def pp_def] by auto
       
  4756         define r1 where "r1=min r e1 / 2"
       
  4757         have "r1<e1" unfolding r1_def using \<open>e1>0\<close> \<open>r>0\<close> by auto
       
  4758         moreover have "r1>0" "pp holomorphic_on cball p r1"
       
  4759                   "(\<forall>w\<in>cball p r1 - {p}. f w = pp w * (w - p) powr of_int po \<and> pp w \<noteq> 0)"
       
  4760           unfolding r1_def using \<open>e1>0\<close> r by auto
       
  4761         ultimately show ?thesis using that \<open>pp p\<noteq>0\<close> by auto
       
  4762       qed
       
  4763 
       
  4764       define e2 where "e2 \<equiv> r/2"
       
  4765       have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto
       
  4766       define anal where "anal \<equiv> \<lambda>w. deriv pp w * h w / pp w"
       
  4767       define prin where "prin \<equiv> \<lambda>w. po * h w / (w - p)"
       
  4768       have "((\<lambda>w.  prin w + anal w) has_contour_integral c * po * h p) (circlepath p e2)"
       
  4769       proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified])
       
  4770         have "ball p r \<subseteq> s"
       
  4771           using \<open>r<e1\<close> avoid_def ball_subset_cball e1_avoid by (simp add: subset_eq)
       
  4772         then have "cball p e2 \<subseteq> s"
       
  4773           using \<open>r>0\<close> unfolding e2_def by auto
       
  4774         then have "(\<lambda>w. po * h w) holomorphic_on cball p e2"
       
  4775           using h_holo by (auto intro!: holomorphic_intros)
       
  4776         then show "(prin has_contour_integral c * po * h p ) (circlepath p e2)"
       
  4777           using Cauchy_integral_circlepath_simple[folded c_def, of "\<lambda>w. po * h w"] \<open>e2>0\<close>
       
  4778           unfolding prin_def by (auto simp add: mult.assoc)
       
  4779         have "anal holomorphic_on ball p r" unfolding anal_def
       
  4780           using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close> \<open>pp p\<noteq>0\<close>
       
  4781           by (auto intro!: holomorphic_intros)
       
  4782         then show "(anal has_contour_integral 0) (circlepath p e2)"
       
  4783           using e2_def \<open>r>0\<close>
       
  4784           by (auto elim!: Cauchy_theorem_disc_simple)
       
  4785       qed
       
  4786       then have "cont ff' p e2" unfolding cont_def po_def
       
  4787       proof (elim has_contour_integral_eq)
       
  4788         fix w assume "w \<in> path_image (circlepath p e2)"
       
  4789         then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
       
  4790         define wp where "wp \<equiv> w-p"
       
  4791         have "wp\<noteq>0" and "pp w \<noteq>0"
       
  4792           unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto
       
  4793         moreover have der_f':"deriv f' w = po * pp w * (w-p) powr (po - 1) + deriv pp w * (w-p) powr po"
       
  4794         proof (rule DERIV_imp_deriv)
       
  4795           have "(pp has_field_derivative (deriv pp w)) (at w)"
       
  4796             using DERIV_deriv_iff_has_field_derivative pp_holo \<open>w\<noteq>p\<close>
       
  4797             by (meson open_ball \<open>w \<in> ball p r\<close> ball_subset_cball holomorphic_derivI holomorphic_on_subset)
       
  4798           then show " (f' has_field_derivative of_int po * pp w * (w - p) powr of_int (po - 1)
       
  4799                   + deriv pp w * (w - p) powr of_int po) (at w)"
       
  4800             unfolding f'_def using \<open>w\<noteq>p\<close>
       
  4801             by (auto intro!: derivative_eq_intros DERIV_cong[OF has_field_derivative_powr_of_int])
       
  4802         qed
       
  4803         ultimately show "prin w + anal w = ff' w"
       
  4804           unfolding ff'_def prin_def anal_def
       
  4805           apply simp
       
  4806           apply (unfold f'_def)
       
  4807           apply (fold wp_def)
       
  4808           apply (auto simp add:field_simps)
       
  4809           by (metis (no_types, lifting) diff_add_cancel mult.commute powr_add powr_to_1)
       
  4810       qed
       
  4811       then have "cont ff p e2" unfolding cont_def
       
  4812       proof (elim has_contour_integral_eq)
       
  4813         fix w assume "w \<in> path_image (circlepath p e2)"
       
  4814         then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
       
  4815         have "deriv f' w =  deriv f w"
       
  4816         proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"])
       
  4817           show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo
       
  4818             by (auto intro!: holomorphic_intros)
       
  4819         next
       
  4820           have "ball p e1 - {p} \<subseteq> s - poles"
       
  4821             using ball_subset_cball e1_avoid[unfolded avoid_def] unfolding pz_def
       
  4822             by auto
       
  4823           then have "ball p r - {p} \<subseteq> s - poles"
       
  4824             apply (elim dual_order.trans)
       
  4825             using \<open>r<e1\<close> by auto
       
  4826           then show "f holomorphic_on ball p r - {p}" using f_holo
       
  4827             by auto
       
  4828         next
       
  4829           show "open (ball p r - {p})" by auto
       
  4830           show "w \<in> ball p r - {p}" using \<open>w\<in>ball p r\<close> \<open>w\<noteq>p\<close> by auto
       
  4831         next
       
  4832           fix x assume "x \<in> ball p r - {p}"
       
  4833           then show "f' x = f x"
       
  4834             using pp_po unfolding f'_def by auto
       
  4835         qed
       
  4836         moreover have " f' w  =  f w "
       
  4837           using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po \<open>w\<noteq>p\<close>
       
  4838           unfolding f'_def by auto
       
  4839         ultimately show "ff' w = ff w"
       
  4840           unfolding ff'_def ff_def by simp
       
  4841       qed
       
  4842       moreover have "cball p e2 \<subseteq> ball p e1"
       
  4843         using \<open>0 < r\<close> \<open>r<e1\<close> e2_def by auto
       
  4844       ultimately show ?thesis using \<open>e2>0\<close> by auto
       
  4845     qed
       
  4846     then obtain e2 where e2:"p\<in>pz \<longrightarrow> e2>0 \<and> cball p e2 \<subseteq> ball p e1 \<and> cont ff p e2"
       
  4847       by auto
       
  4848     define e4 where "e4 \<equiv> if p\<in>pz then e2 else  e1"
       
  4849     have "e4>0" using e2 \<open>e1>0\<close> unfolding e4_def by auto
       
  4850     moreover have "avoid p e4" using e2 \<open>e1>0\<close> e1_avoid unfolding e4_def avoid_def by auto
       
  4851     moreover have "p\<in>pz \<longrightarrow> cont ff p e4"
       
  4852       by (auto simp add: e2 e4_def)
       
  4853     ultimately show ?thesis by auto
       
  4854   qed
       
  4855   then obtain get_e where get_e:"\<forall>p\<in>s. get_e p>0 \<and> avoid p (get_e p)
       
  4856       \<and> (p\<in>pz \<longrightarrow> cont ff p (get_e p))"
       
  4857     by metis
       
  4858   define ci where "ci \<equiv> \<lambda>p. contour_integral (circlepath p (get_e p)) ff"
       
  4859   define w where "w \<equiv> \<lambda>p. winding_number g p"
       
  4860   have "contour_integral g ff = (\<Sum>p\<in>pz. w p * ci p)" unfolding ci_def w_def
       
  4861   proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected s\<close> finite _ \<open>valid_path g\<close> loop
       
  4862         path_img homo])
       
  4863     have "open (s - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto
       
  4864     then show "ff holomorphic_on s - pz" unfolding ff_def using f_holo h_holo
       
  4865       by (auto intro!: holomorphic_intros simp add:pz_def)
       
  4866   next
       
  4867     show "\<forall>p\<in>s. 0 < get_e p \<and> (\<forall>w\<in>cball p (get_e p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pz))"
       
  4868       using get_e using avoid_def by blast
       
  4869   qed
       
  4870   also have "... = (\<Sum>p\<in>pz. c * w p * h p * zorder f p)"
       
  4871   proof (rule sum.cong[of pz pz,simplified])
       
  4872     fix p assume "p \<in> pz"
       
  4873     show "w p * ci p = c * w p * h p * (zorder f p)"
       
  4874     proof (cases "p\<in>s")
       
  4875       assume "p \<in> s"
       
  4876       have "ci p = c * h p * (zorder f p)" unfolding ci_def
       
  4877         apply (rule contour_integral_unique)
       
  4878         using get_e \<open>p\<in>s\<close> \<open>p\<in>pz\<close> unfolding cont_def by (metis mult.assoc mult.commute)
       
  4879       thus ?thesis by auto
       
  4880     next
       
  4881       assume "p\<notin>s"
       
  4882       then have "w p=0" using homo unfolding w_def by auto
       
  4883       then show ?thesis by auto
       
  4884     qed
       
  4885   qed
       
  4886   also have "... = c*(\<Sum>p\<in>pz. w p * h p * zorder f p)"
       
  4887     unfolding sum_distrib_left by (simp add:algebra_simps)
       
  4888   finally have "contour_integral g ff = c * (\<Sum>p\<in>pz. w p * h p * of_int (zorder f p))" .
       
  4889   then show ?thesis unfolding ff_def c_def w_def by simp
       
  4890 qed
       
  4891 
       
  4892 subsection \<open>Rouche's theorem \<close>
       
  4893 
       
  4894 theorem Rouche_theorem:
       
  4895   fixes f g::"complex \<Rightarrow> complex" and s:: "complex set"
       
  4896   defines "fg\<equiv>(\<lambda>p. f p + g p)"
       
  4897   defines "zeros_fg\<equiv>{p. fg p = 0}" and "zeros_f\<equiv>{p. f p = 0}"
       
  4898   assumes
       
  4899     "open s" and "connected s" and
       
  4900     "finite zeros_fg" and
       
  4901     "finite zeros_f" and
       
  4902     f_holo:"f holomorphic_on s" and
       
  4903     g_holo:"g holomorphic_on s" and
       
  4904     "valid_path \<gamma>" and
       
  4905     loop:"pathfinish \<gamma> = pathstart \<gamma>" and
       
  4906     path_img:"path_image \<gamma> \<subseteq> s " and
       
  4907     path_less:"\<forall>z\<in>path_image \<gamma>. cmod(f z) > cmod(g z)" and
       
  4908     homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number \<gamma> z = 0"
       
  4909   shows "(\<Sum>p\<in>zeros_fg. winding_number \<gamma> p * zorder fg p)
       
  4910           = (\<Sum>p\<in>zeros_f. winding_number \<gamma> p * zorder f p)"
       
  4911 proof -
       
  4912   have path_fg:"path_image \<gamma> \<subseteq> s - zeros_fg"
       
  4913   proof -
       
  4914     have False when "z\<in>path_image \<gamma>" and "f z + g z=0" for z
       
  4915     proof -
       
  4916       have "cmod (f z) > cmod (g z)" using \<open>z\<in>path_image \<gamma>\<close> path_less by auto
       
  4917       moreover have "f z = - g z"  using \<open>f z + g z =0\<close> by (simp add: eq_neg_iff_add_eq_0)
       
  4918       then have "cmod (f z) = cmod (g z)" by auto
       
  4919       ultimately show False by auto
       
  4920     qed
       
  4921     then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto
       
  4922   qed
       
  4923   have path_f:"path_image \<gamma> \<subseteq> s - zeros_f"
       
  4924   proof -
       
  4925     have False when "z\<in>path_image \<gamma>" and "f z =0" for z
       
  4926     proof -
       
  4927       have "cmod (g z) < cmod (f z) " using \<open>z\<in>path_image \<gamma>\<close> path_less by auto
       
  4928       then have "cmod (g z) < 0" using \<open>f z=0\<close> by auto
       
  4929       then show False by auto
       
  4930     qed
       
  4931     then show ?thesis unfolding zeros_f_def using path_img by auto
       
  4932   qed
       
  4933   define w where "w \<equiv> \<lambda>p. winding_number \<gamma> p"
       
  4934   define c where "c \<equiv> 2 * complex_of_real pi * \<i>"
       
  4935   define h where "h \<equiv> \<lambda>p. g p / f p + 1"
       
  4936   obtain spikes
       
  4937     where "finite spikes" and spikes: "\<forall>x\<in>{0..1} - spikes. \<gamma> differentiable at x"
       
  4938     using \<open>valid_path \<gamma>\<close>
       
  4939     by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
  4940   have h_contour:"((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>"
       
  4941   proof -
       
  4942     have outside_img:"0 \<in> outside (path_image (h o \<gamma>))"
       
  4943     proof -
       
  4944       have "h p \<in> ball 1 1" when "p\<in>path_image \<gamma>" for p
       
  4945       proof -
       
  4946         have "cmod (g p/f p) <1" using path_less[rule_format,OF that]
       
  4947           apply (cases "cmod (f p) = 0")
       
  4948           by (auto simp add: norm_divide)
       
  4949         then show ?thesis unfolding h_def by (auto simp add:dist_complex_def)
       
  4950       qed
       
  4951       then have "path_image (h o \<gamma>) \<subseteq> ball 1 1"
       
  4952         by (simp add: image_subset_iff path_image_compose)
       
  4953       moreover have " (0::complex) \<notin> ball 1 1" by (simp add: dist_norm)
       
  4954       ultimately show "?thesis"
       
  4955         using  convex_in_outside[of "ball 1 1" 0] outside_mono by blast
       
  4956     qed
       
  4957     have valid_h:"valid_path (h \<circ> \<gamma>)"
       
  4958     proof (rule valid_path_compose_holomorphic[OF \<open>valid_path \<gamma>\<close> _ _ path_f])
       
  4959       show "h holomorphic_on s - zeros_f"
       
  4960         unfolding h_def using f_holo g_holo
       
  4961         by (auto intro!: holomorphic_intros simp add:zeros_f_def)
       
  4962     next
       
  4963       show "open (s - zeros_f)" using \<open>finite zeros_f\<close> \<open>open s\<close> finite_imp_closed
       
  4964         by auto
       
  4965     qed
       
  4966     have "((\<lambda>z. 1/z) has_contour_integral 0) (h \<circ> \<gamma>)"
       
  4967     proof -
       
  4968       have "0 \<notin> path_image (h \<circ> \<gamma>)" using outside_img by (simp add: outside_def)
       
  4969       then have "((\<lambda>z. 1/z) has_contour_integral c * winding_number (h \<circ> \<gamma>) 0) (h \<circ> \<gamma>)"
       
  4970         using has_contour_integral_winding_number[of "h o \<gamma>" 0,simplified] valid_h
       
  4971         unfolding c_def by auto
       
  4972       moreover have "winding_number (h o \<gamma>) 0 = 0"
       
  4973       proof -
       
  4974         have "0 \<in> outside (path_image (h \<circ> \<gamma>))" using outside_img .
       
  4975         moreover have "path (h o \<gamma>)"
       
  4976           using valid_h  by (simp add: valid_path_imp_path)
       
  4977         moreover have "pathfinish (h o \<gamma>) = pathstart (h o \<gamma>)"
       
  4978           by (simp add: loop pathfinish_compose pathstart_compose)
       
  4979         ultimately show ?thesis using winding_number_zero_in_outside by auto
       
  4980       qed
       
  4981       ultimately show ?thesis by auto
       
  4982     qed
       
  4983     moreover have "vector_derivative (h \<circ> \<gamma>) (at x) = vector_derivative \<gamma> (at x) * deriv h (\<gamma> x)"
       
  4984       when "x\<in>{0..1} - spikes" for x
       
  4985     proof (rule vector_derivative_chain_at_general)
       
  4986       show "\<gamma> differentiable at x" using that \<open>valid_path \<gamma>\<close> spikes by auto
       
  4987     next
       
  4988       define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
       
  4989       define t where "t \<equiv> \<gamma> x"
       
  4990       have "f t\<noteq>0" unfolding zeros_f_def t_def
       
  4991         by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that)
       
  4992       moreover have "t\<in>s"
       
  4993         using contra_subsetD path_image_def path_fg t_def that by fastforce
       
  4994       ultimately have "(h has_field_derivative der t) (at t)"
       
  4995         unfolding h_def der_def using g_holo f_holo \<open>open s\<close>
       
  4996         by (auto intro!: holomorphic_derivI derivative_eq_intros)
       
  4997       then show "h field_differentiable at (\<gamma> x)"
       
  4998         unfolding t_def field_differentiable_def by blast
       
  4999     qed
       
  5000     then have " ((/) 1 has_contour_integral 0) (h \<circ> \<gamma>)
       
  5001                   = ((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>"
       
  5002       unfolding has_contour_integral
       
  5003       apply (intro has_integral_spike_eq[OF negligible_finite, OF \<open>finite spikes\<close>])
       
  5004       by auto
       
  5005     ultimately show ?thesis by auto
       
  5006   qed
       
  5007   then have "contour_integral \<gamma> (\<lambda>x. deriv h x / h x) = 0"
       
  5008     using  contour_integral_unique by simp
       
  5009   moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = contour_integral \<gamma> (\<lambda>x. deriv f x / f x)
       
  5010       + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)"
       
  5011   proof -
       
  5012     have "(\<lambda>p. deriv f p / f p) contour_integrable_on \<gamma>"
       
  5013     proof (rule contour_integrable_holomorphic_simple[OF _ _ \<open>valid_path \<gamma>\<close> path_f])
       
  5014       show "open (s - zeros_f)" using finite_imp_closed[OF \<open>finite zeros_f\<close>] \<open>open s\<close>
       
  5015         by auto
       
  5016       then show "(\<lambda>p. deriv f p / f p) holomorphic_on s - zeros_f"
       
  5017         using f_holo
       
  5018         by (auto intro!: holomorphic_intros simp add:zeros_f_def)
       
  5019     qed
       
  5020     moreover have "(\<lambda>p. deriv h p / h p) contour_integrable_on \<gamma>"
       
  5021       using h_contour
       
  5022       by (simp add: has_contour_integral_integrable)
       
  5023     ultimately have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x + deriv h x / h x) =
       
  5024                         contour_integral \<gamma> (\<lambda>p. deriv f p / f p) + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)"
       
  5025       using contour_integral_add[of "(\<lambda>p. deriv f p / f p)" \<gamma> "(\<lambda>p. deriv h p / h p)" ]
       
  5026       by auto
       
  5027     moreover have "deriv fg p / fg p =  deriv f p / f p + deriv h p / h p"
       
  5028                       when "p\<in> path_image \<gamma>" for p
       
  5029     proof -
       
  5030       have "fg p\<noteq>0" and "f p\<noteq>0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def
       
  5031         by auto
       
  5032       have "h p\<noteq>0"
       
  5033       proof (rule ccontr)
       
  5034         assume "\<not> h p \<noteq> 0"
       
  5035         then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2)
       
  5036         then have "cmod (g p/f p) = 1" by auto
       
  5037         moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that]
       
  5038           apply (cases "cmod (f p) = 0")
       
  5039           by (auto simp add: norm_divide)
       
  5040         ultimately show False by auto
       
  5041       qed
       
  5042       have der_fg:"deriv fg p =  deriv f p + deriv g p" unfolding fg_def
       
  5043         using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _  \<open>open s\<close>] path_img that
       
  5044         by auto
       
  5045       have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)"
       
  5046       proof -
       
  5047         define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
       
  5048         have "p\<in>s" using path_img that by auto
       
  5049         then have "(h has_field_derivative der p) (at p)"
       
  5050           unfolding h_def der_def using g_holo f_holo \<open>open s\<close> \<open>f p\<noteq>0\<close>
       
  5051           by (auto intro!: derivative_eq_intros holomorphic_derivI)
       
  5052         then show ?thesis unfolding der_def using DERIV_imp_deriv by auto
       
  5053       qed
       
  5054       show ?thesis
       
  5055         apply (simp only:der_fg der_h)
       
  5056         apply (auto simp add:field_simps \<open>h p\<noteq>0\<close> \<open>f p\<noteq>0\<close> \<open>fg p\<noteq>0\<close>)
       
  5057         by (auto simp add:field_simps h_def \<open>f p\<noteq>0\<close> fg_def)
       
  5058     qed
       
  5059     then have "contour_integral \<gamma> (\<lambda>p. deriv fg p / fg p)
       
  5060                   = contour_integral \<gamma> (\<lambda>p. deriv f p / f p + deriv h p / h p)"
       
  5061       by (elim contour_integral_eq)
       
  5062     ultimately show ?thesis by auto
       
  5063   qed
       
  5064   moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = c * (\<Sum>p\<in>zeros_fg. w p * zorder fg p)"
       
  5065     unfolding c_def zeros_fg_def w_def
       
  5066   proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo
       
  5067         , of _ "{}" "\<lambda>_. 1",simplified])
       
  5068     show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto
       
  5069     show "path_image \<gamma> \<subseteq> s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def .
       
  5070     show " finite {p. fg p = 0}" using \<open>finite zeros_fg\<close> unfolding zeros_fg_def .
       
  5071   qed
       
  5072   moreover have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x) = c * (\<Sum>p\<in>zeros_f. w p * zorder f p)"
       
  5073     unfolding c_def zeros_f_def w_def
       
  5074   proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo
       
  5075         , of _ "{}" "\<lambda>_. 1",simplified])
       
  5076     show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto
       
  5077     show "path_image \<gamma> \<subseteq> s - {p. f p = 0}" using path_f unfolding zeros_f_def .
       
  5078     show " finite {p. f p = 0}" using \<open>finite zeros_f\<close> unfolding zeros_f_def .
       
  5079   qed
       
  5080   ultimately have " c* (\<Sum>p\<in>zeros_fg. w p * (zorder fg p)) = c* (\<Sum>p\<in>zeros_f. w p * (zorder f p))"
       
  5081     by auto
       
  5082   then show ?thesis unfolding c_def using w_def by auto
       
  5083 qed
       
  5084 
       
  5085 end