Theory Conformal_Mappings

theory Conformal_Mappings
imports Cauchy_Integral_Theorem
section ‹Conformal Mappings. Consequences of Cauchy's integral theorem.›

text‹By John Harrison et al.  Ported from HOL Light by L C Paulson (2016)›

text‹Also Cauchy's residue theorem by Wenda Li (2016)›

theory Conformal_Mappings
imports Cauchy_Integral_Theorem

begin

subsection‹Cauchy's inequality and more versions of Liouville›

lemma Cauchy_higher_deriv_bound:
    assumes holf: "f holomorphic_on (ball z r)"
        and contf: "continuous_on (cball z r) f"
        and fin : "⋀w. w ∈ ball z r ⟹ f w ∈ ball y B0"
        and "0 < r" and "0 < n"
      shows "norm ((deriv ^^ n) f z) ≤ (fact n) * B0 / r^n"
proof -
  have "0 < B0" using ‹0 < r› fin [of z]
    by (metis ball_eq_empty ex_in_conv fin not_less)
  have le_B0: "⋀w. cmod (w - z) ≤ r ⟹ cmod (f w - y) ≤ B0"
    apply (rule continuous_on_closure_norm_le [of "ball z r" "λw. f w - y"])
    apply (auto simp: ‹0 < r›  dist_norm norm_minus_commute)
    apply (rule continuous_intros contf)+
    using fin apply (simp add: dist_commute dist_norm less_eq_real_def)
    done
  have "(deriv ^^ n) f z = (deriv ^^ n) (λw. f w) z - (deriv ^^ n) (λw. y) z"
    using ‹0 < n› by simp
  also have "... = (deriv ^^ n) (λw. f w - y) z"
    by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: ‹0 < r›)
  finally have "(deriv ^^ n) f z = (deriv ^^ n) (λw. f w - y) z" .
  have contf': "continuous_on (cball z r) (λu. f u - y)"
    by (rule contf continuous_intros)+
  have holf': "(λu. (f u - y)) holomorphic_on (ball z r)"
    by (simp add: holf holomorphic_on_diff)
  define a where "a = (2 * pi)/(fact n)"
  have "0 < a"  by (simp add: a_def)
  have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)"
    using ‹0 < r› by (simp add: a_def divide_simps)
  have der_dif: "(deriv ^^ n) (λw. f w - y) z = (deriv ^^ n) f z"
    using ‹0 < r› ‹0 < n›
    by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const])
  have "norm ((2 * of_real pi * 𝗂)/(fact n) * (deriv ^^ n) (λw. f w - y) z)
        ≤ (B0/r^(Suc n)) * (2 * pi * r)"
    apply (rule has_contour_integral_bound_circlepath [of "(λu. (f u - y)/(u - z)^(Suc n))" _ z])
    using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf']
    using ‹0 < B0› ‹0 < r›
    apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0)
    done
  then show ?thesis
    using ‹0 < r›
    by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0)
qed

proposition Cauchy_inequality:
    assumes holf: "f holomorphic_on (ball ξ r)"
        and contf: "continuous_on (cball ξ r) f"
        and "0 < r"
        and nof: "⋀x. norm(ξ-x) = r ⟹ norm(f x) ≤ B"
      shows "norm ((deriv ^^ n) f ξ) ≤ (fact n) * B / r^n"
proof -
  obtain x where "norm (ξ-x) = r"
    by (metis abs_of_nonneg add_diff_cancel_left' ‹0 < r› diff_add_cancel
                 dual_order.strict_implies_order norm_of_real)
  then have "0 ≤ B"
    by (metis nof norm_not_less_zero not_le order_trans)
  have  "((λu. f u / (u - ξ) ^ Suc n) has_contour_integral (2 * pi) * 𝗂 / fact n * (deriv ^^ n) f ξ)
         (circlepath ξ r)"
    apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf])
    using ‹0 < r› by simp
  then have "norm ((2 * pi * 𝗂)/(fact n) * (deriv ^^ n) f ξ) ≤ (B / r^(Suc n)) * (2 * pi * r)"
    apply (rule has_contour_integral_bound_circlepath)
    using ‹0 ≤ B› ‹0 < r›
    apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc)
    done
  then show ?thesis using ‹0 < r›
    by (simp add: norm_divide norm_mult field_simps)
qed

proposition Liouville_polynomial:
    assumes holf: "f holomorphic_on UNIV"
        and nof: "⋀z. A ≤ norm z ⟹ norm(f z) ≤ B * norm z ^ n"
      shows "f ξ = (∑k≤n. (deriv^^k) f 0 / fact k * ξ ^ k)"
proof (cases rule: le_less_linear [THEN disjE])
  assume "B ≤ 0"
  then have "⋀z. A ≤ norm z ⟹ norm(f z) = 0"
    by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le)
  then have f0: "(f ⤏ 0) at_infinity"
    using Lim_at_infinity by force
  then have [simp]: "f = (λw. 0)"
    using Liouville_weak [OF holf, of 0]
    by (simp add: eventually_at_infinity f0) meson
  show ?thesis by simp
next
  assume "0 < B"
  have "((λk. (deriv ^^ k) f 0 / (fact k) * (ξ - 0)^k) sums f ξ)"
    apply (rule holomorphic_power_series [where r = "norm ξ + 1"])
    using holf holomorphic_on_subset apply auto
    done
  then have sumsf: "((λk. (deriv ^^ k) f 0 / (fact k) * ξ^k) sums f ξ)" by simp
  have "(deriv ^^ k) f 0 / fact k * ξ ^ k = 0" if "k>n" for k
  proof (cases "(deriv ^^ k) f 0 = 0")
    case True then show ?thesis by simp
  next
    case False
    define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (¦A¦ + 1))"
    have "1 ≤ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (¦A¦ + 1))"
      using ‹0 < B› by simp
    then have wge1: "1 ≤ norm w"
      by (metis norm_of_real w_def)
    then have "w ≠ 0" by auto
    have kB: "0 < fact k * B"
      using ‹0 < B› by simp
    then have "0 ≤ fact k * B / cmod ((deriv ^^ k) f 0)"
      by simp
    then have wgeA: "A ≤ cmod w"
      by (simp only: w_def norm_of_real)
    have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (¦A¦ + 1))"
      using ‹0 < B› by simp
    then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w"
      by (metis norm_of_real w_def)
    then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)"
      using False by (simp add: divide_simps mult.commute split: if_split_asm)
    also have "... ≤ fact k * (B * norm w ^ n) / norm w ^ k"
      apply (rule Cauchy_inequality)
         using holf holomorphic_on_subset apply force
        using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast
       using ‹w ≠ 0› apply (simp add:)
       by (metis nof wgeA dist_0_norm dist_norm)
    also have "... = fact k * (B * 1 / cmod w ^ (k-n))"
      apply (simp only: mult_cancel_left times_divide_eq_right [symmetric])
      using ‹k>n› ‹w ≠ 0› ‹0 < B› apply (simp add: divide_simps semiring_normalization_rules)
      done
    also have "... = fact k * B / cmod w ^ (k-n)"
      by simp
    finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" .
    then have "1 / cmod w < 1 / cmod w ^ (k - n)"
      by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos)
    then have "cmod w ^ (k - n) < cmod w"
      by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one)
    with self_le_power [OF wge1] have False
      by (meson diff_is_0_eq not_gr0 not_le that)
    then show ?thesis by blast
  qed
  then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * ξ ^ (k + Suc n) = 0" for k
    using not_less_eq by blast
  then have "(λi. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * ξ ^ (i + Suc n)) sums 0"
    by (rule sums_0)
  with sums_split_initial_segment [OF sumsf, where n = "Suc n"]
  show ?thesis
    using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce
qed

text‹Every bounded entire function is a constant function.›
theorem Liouville_theorem:
    assumes holf: "f holomorphic_on UNIV"
        and bf: "bounded (range f)"
    obtains c where "⋀z. f z = c"
proof -
  obtain B where "⋀z. cmod (f z) ≤ B"
    by (meson bf bounded_pos rangeI)
  then show ?thesis
    using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast
qed



text‹A holomorphic function f has only isolated zeros unless f is 0.›

proposition powser_0_nonzero:
  fixes a :: "nat ⇒ 'a::{real_normed_field,banach}"
  assumes r: "0 < r"
      and sm: "⋀x. norm (x - ξ) < r ⟹ (λn. a n * (x - ξ) ^ n) sums (f x)"
      and [simp]: "f ξ = 0"
      and m0: "a m ≠ 0" and "m>0"
  obtains s where "0 < s" and "⋀z. z ∈ cball ξ s - {ξ} ⟹ f z ≠ 0"
proof -
  have "r ≤ conv_radius a"
    using sm sums_summable by (auto simp: le_conv_radius_iff [where ξ=ξ])
  obtain m where am: "a m ≠ 0" and az [simp]: "(⋀n. n<m ⟹ a n = 0)"
    apply (rule_tac m = "LEAST n. a n ≠ 0" in that)
    using m0
    apply (rule LeastI2)
    apply (fastforce intro:  dest!: not_less_Least)+
    done
  define b where "b i = a (i+m) / a m" for i
  define g where "g x = suminf (λi. b i * (x - ξ) ^ i)" for x
  have [simp]: "b 0 = 1"
    by (simp add: am b_def)
  { fix x::'a
    assume "norm (x - ξ) < r"
    then have "(λn. (a m * (x - ξ)^m) * (b n * (x - ξ)^n)) sums (f x)"
      using am az sm sums_zero_iff_shift [of m "(λn. a n * (x - ξ) ^ n)" "f x"]
      by (simp add: b_def monoid_mult_class.power_add algebra_simps)
    then have "x ≠ ξ ⟹ (λn. b n * (x - ξ)^n) sums (f x / (a m * (x - ξ)^m))"
      using am by (simp add: sums_mult_D)
  } note bsums = this
  then have  "norm (x - ξ) < r ⟹ summable (λn. b n * (x - ξ)^n)" for x
    using sums_summable by (cases "x=ξ") auto
  then have "r ≤ conv_radius b"
    by (simp add: le_conv_radius_iff [where ξ=ξ])
  then have "r/2 < conv_radius b"
    using not_le order_trans r by fastforce
  then have "continuous_on (cball ξ (r/2)) g"
    using powser_continuous_suminf [of "r/2" b ξ] by (simp add: g_def)
  then obtain s where "s>0"  "⋀x. ⟦norm (x - ξ) ≤ s; norm (x - ξ) ≤ r/2⟧ ⟹ dist (g x) (g ξ) < 1/2"
    apply (rule continuous_onE [where x=ξ and e = "1/2"])
    using r apply (auto simp: norm_minus_commute dist_norm)
    done
  moreover have "g ξ = 1"
    by (simp add: g_def)
  ultimately have gnz: "⋀x. ⟦norm (x - ξ) ≤ s; norm (x - ξ) ≤ r/2⟧ ⟹ (g x) ≠ 0"
    by fastforce
  have "f x ≠ 0" if "x ≠ ξ" "norm (x - ξ) ≤ s" "norm (x - ξ) ≤ r/2" for x
    using bsums [of x] that gnz [of x]
    apply (auto simp: g_def)
    using r sums_iff by fastforce
  then show ?thesis
    apply (rule_tac s="min s (r/2)" in that)
    using ‹0 < r› ‹0 < s› by (auto simp: dist_commute dist_norm)
qed

proposition isolated_zeros:
  assumes holf: "f holomorphic_on S"
      and "open S" "connected S" "ξ ∈ S" "f ξ = 0" "β ∈ S" "f β ≠ 0"
  obtains r where "0 < r" "ball ξ r ⊆ S" "⋀z. z ∈ ball ξ r - {ξ} ⟹ f z ≠ 0"
proof -
  obtain r where "0 < r" and r: "ball ξ r ⊆ S"
    using ‹open S› ‹ξ ∈ S› open_contains_ball_eq by blast
  have powf: "((λn. (deriv ^^ n) f ξ / (fact n) * (z - ξ)^n) sums f z)" if "z ∈ ball ξ r" for z
    apply (rule holomorphic_power_series [OF _ that])
    apply (rule holomorphic_on_subset [OF holf r])
    done
  obtain m where m: "(deriv ^^ m) f ξ / (fact m) ≠ 0"
    using holomorphic_fun_eq_0_on_connected [OF holf ‹open S› ‹connected S› _ ‹ξ ∈ S› ‹β ∈ S›] ‹f β ≠ 0›
    by auto
  then have "m ≠ 0" using assms(5) funpow_0 by fastforce
  obtain s where "0 < s" and s: "⋀z. z ∈ cball ξ s - {ξ} ⟹ f z ≠ 0"
    apply (rule powser_0_nonzero [OF ‹0 < r› powf ‹f ξ = 0› m])
    using ‹m ≠ 0› by (auto simp: dist_commute dist_norm)
  have "0 < min r s"  by (simp add: ‹0 < r› ‹0 < s›)
  then show ?thesis
    apply (rule that)
    using r s by auto
qed


proposition analytic_continuation:
  assumes holf: "f holomorphic_on S"
      and S: "open S" "connected S"
      and "U ⊆ S" "ξ ∈ S"
      and "ξ islimpt U"
      and fU0 [simp]: "⋀z. z ∈ U ⟹ f z = 0"
      and "w ∈ S"
    shows "f w = 0"
proof -
  obtain e where "0 < e" and e: "cball ξ e ⊆ S"
    using ‹open S› ‹ξ ∈ S› open_contains_cball_eq by blast
  define T where "T = cball ξ e ∩ U"
  have contf: "continuous_on (closure T) f"
    by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on
              holomorphic_on_subset inf.cobounded1)
  have fT0 [simp]: "⋀x. x ∈ T ⟹ f x = 0"
    by (simp add: T_def)
  have "⋀r. ⟦∀e>0. ∃x'∈U. x' ≠ ξ ∧ dist x' ξ < e; 0 < r⟧ ⟹ ∃x'∈cball ξ e ∩ U. x' ≠ ξ ∧ dist x' ξ < r"
    by (metis ‹0 < e› IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj)
  then have "ξ islimpt T" using ‹ξ islimpt U›
    by (auto simp: T_def islimpt_approachable)
  then have "ξ ∈ closure T"
    by (simp add: closure_def)
  then have "f ξ = 0"
    by (auto simp: continuous_constant_on_closure [OF contf])
  show ?thesis
    apply (rule ccontr)
    apply (rule isolated_zeros [OF holf ‹open S› ‹connected S› ‹ξ ∈ S› ‹f ξ = 0› ‹w ∈ S›], assumption)
    by (metis open_ball ‹ξ islimpt T› centre_in_ball fT0 insertE insert_Diff islimptE)
qed

corollary analytic_continuation_open:
  assumes "open s" "open s'" "s ≠ {}" "connected s'" "s ⊆ s'"
  assumes "f holomorphic_on s'" "g holomorphic_on s'" "⋀z. z ∈ s ⟹ f z = g z"
  assumes "z ∈ s'"
  shows   "f z = g z"
proof -
  from ‹s ≠ {}› obtain ξ where "ξ ∈ s" by auto
  with ‹open s› have ξ: "ξ islimpt s" 
    by (intro interior_limit_point) (auto simp: interior_open)
  have "f z - g z = 0"
    by (rule analytic_continuation[of "λz. f z - g z" s' s ξ])
       (insert assms ‹ξ ∈ s› ξ, auto intro: holomorphic_intros)
  thus ?thesis by simp
qed


subsection‹Open mapping theorem›

lemma holomorphic_contract_to_zero:
  assumes contf: "continuous_on (cball ξ r) f"
      and holf: "f holomorphic_on ball ξ r"
      and "0 < r"
      and norm_less: "⋀z. norm(ξ - z) = r ⟹ norm(f ξ) < norm(f z)"
  obtains z where "z ∈ ball ξ r" "f z = 0"
proof -
  { assume fnz: "⋀w. w ∈ ball ξ r ⟹ f w ≠ 0"
    then have "0 < norm (f ξ)"
      by (simp add: ‹0 < r›)
    have fnz': "⋀w. w ∈ cball ξ r ⟹ f w ≠ 0"
      by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero)
    have "frontier(cball ξ r) ≠ {}"
      using ‹0 < r› by simp
    define g where [abs_def]: "g z = inverse (f z)" for z
    have contg: "continuous_on (cball ξ r) g"
      unfolding g_def using contf continuous_on_inverse fnz' by blast
    have holg: "g holomorphic_on ball ξ r"
      unfolding g_def using fnz holf holomorphic_on_inverse by blast
    have "frontier (cball ξ r) ⊆ cball ξ r"
      by (simp add: subset_iff)
    then have contf': "continuous_on (frontier (cball ξ r)) f"
          and contg': "continuous_on (frontier (cball ξ r)) g"
      by (blast intro: contf contg continuous_on_subset)+
    have froc: "frontier(cball ξ r) ≠ {}"
      using ‹0 < r› by simp
    moreover have "continuous_on (frontier (cball ξ r)) (norm o f)"
      using contf' continuous_on_compose continuous_on_norm_id by blast
    ultimately obtain w where w: "w ∈ frontier(cball ξ r)"
                          and now: "⋀x. x ∈ frontier(cball ξ r) ⟹ norm (f w) ≤ norm (f x)"
      apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]])
      apply (simp add:)
      done
    then have fw: "0 < norm (f w)"
      by (simp add: fnz')
    have "continuous_on (frontier (cball ξ r)) (norm o g)"
      using contg' continuous_on_compose continuous_on_norm_id by blast
    then obtain v where v: "v ∈ frontier(cball ξ r)"
               and nov: "⋀x. x ∈ frontier(cball ξ r) ⟹ norm (g v) ≥ norm (g x)"
      apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]])
      apply (simp add:)
      done
    then have fv: "0 < norm (f v)"
      by (simp add: fnz')
    have "norm ((deriv ^^ 0) g ξ) ≤ fact 0 * norm (g v) / r ^ 0"
      by (rule Cauchy_inequality [OF holg contg ‹0 < r›]) (simp add: dist_norm nov)
    then have "cmod (g ξ) ≤ norm (g v)"
      by simp
    with w have wr: "norm (ξ - w) = r" and nfw: "norm (f w) ≤ norm (f ξ)"
      apply (simp_all add: dist_norm)
      by (metis ‹0 < cmod (f ξ)› g_def less_imp_inverse_less norm_inverse not_le now order_trans v)
    with fw have False
      using norm_less by force
  }
  with that show ?thesis by blast
qed


theorem open_mapping_thm:
  assumes holf: "f holomorphic_on S"
      and S: "open S" "connected S"
      and "open U" "U ⊆ S"
      and fne: "~ f constant_on S"
    shows "open (f ` U)"
proof -
  have *: "open (f ` U)"
          if "U ≠ {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "⋀x. ∃y ∈ U. f y ≠ x"
          for U
  proof (clarsimp simp: open_contains_ball)
    fix ξ assume ξ: "ξ ∈ U"
    show "∃e>0. ball (f ξ) e ⊆ f ` U"
    proof -
      have hol: "(λz. f z - f ξ) holomorphic_on U"
        by (rule holomorphic_intros that)+
      obtain s where "0 < s" and sbU: "ball ξ s ⊆ U"
                 and sne: "⋀z. z ∈ ball ξ s - {ξ} ⟹ (λz. f z - f ξ) z ≠ 0"
        using isolated_zeros [OF hol U ξ]  by (metis fneU right_minus_eq)
      obtain r where "0 < r" and r: "cball ξ r ⊆ ball ξ s"
        apply (rule_tac r="s/2" in that)
        using ‹0 < s› by auto
      have "cball ξ r ⊆ U"
        using sbU r by blast
      then have frsbU: "frontier (cball ξ r) ⊆ U"
        using Diff_subset frontier_def order_trans by fastforce
      then have cof: "compact (frontier(cball ξ r))"
        by blast
      have frne: "frontier (cball ξ r) ≠ {}"
        using ‹0 < r› by auto
      have contfr: "continuous_on (frontier (cball ξ r)) (λz. norm (f z - f ξ))"
        apply (rule continuous_on_compose2 [OF Complex_Analysis_Basics.continuous_on_norm_id])
        using hol frsbU holomorphic_on_imp_continuous_on holomorphic_on_subset by blast+
      obtain w where "norm (ξ - w) = r"
                 and w: "(⋀z. norm (ξ - z) = r ⟹ norm (f w - f ξ) ≤ norm(f z - f ξ))"
        apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]])
        apply (simp add: dist_norm)
        done
      moreover define ε where "ε ≡ norm (f w - f ξ) / 3"
      ultimately have "0 < ε"
        using ‹0 < r› dist_complex_def r sne by auto
      have "ball (f ξ) ε ⊆ f ` U"
      proof
        fix γ
        assume γ: "γ ∈ ball (f ξ) ε"
        have *: "cmod (γ - f ξ) < cmod (γ - f z)" if "cmod (ξ - z) = r" for z
        proof -
          have lt: "cmod (f w - f ξ) / 3 < cmod (γ - f z)"
            using w [OF that] γ
            using dist_triangle2 [of "f ξ" "γ"  "f z"] dist_triangle2 [of "f ξ" "f z" γ]
            by (simp add: ε_def dist_norm norm_minus_commute)
          show ?thesis
            by (metis ε_def dist_commute dist_norm less_trans lt mem_ball γ)
       qed
       have "continuous_on (cball ξ r) (λz. γ - f z)"
          apply (rule continuous_intros)+
          using ‹cball ξ r ⊆ U› ‹f holomorphic_on U›
          apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on)
          done
        moreover have "(λz. γ - f z) holomorphic_on ball ξ r"
          apply (rule holomorphic_intros)+
          apply (metis ‹cball ξ r ⊆ U› ‹f holomorphic_on U› holomorphic_on_subset interior_cball interior_subset)
          done
        ultimately obtain z where "z ∈ ball ξ r" "γ - f z = 0"
          apply (rule holomorphic_contract_to_zero)
          apply (blast intro!: ‹0 < r› *)+
          done
        then show "γ ∈ f ` U"
          using ‹cball ξ r ⊆ U› by fastforce
      qed
      then show ?thesis using  ‹0 < ε› by blast
    qed
  qed
  have "open (f ` X)" if "X ∈ components U" for X
  proof -
    have holfU: "f holomorphic_on U"
      using ‹U ⊆ S› holf holomorphic_on_subset by blast
    have "X ≠ {}"
      using that by (simp add: in_components_nonempty)
    moreover have "open X"
      using that ‹open U› open_components by auto
    moreover have "connected X"
      using that in_components_maximal by blast
    moreover have "f holomorphic_on X"
      by (meson that holfU holomorphic_on_subset in_components_maximal)
    moreover have "∃y∈X. f y ≠ x" for x
    proof (rule ccontr)
      assume not: "¬ (∃y∈X. f y ≠ x)"
      have "X ⊆ S"
        using ‹U ⊆ S› in_components_subset that by blast
      obtain w where w: "w ∈ X" using ‹X ≠ {}› by blast
      have wis: "w islimpt X"
        using w ‹open X› interior_eq by auto
      have hol: "(λz. f z - x) holomorphic_on S"
        by (simp add: holf holomorphic_on_diff)
      with fne [unfolded constant_on_def] analytic_continuation [OF hol S ‹X ⊆ S› _ wis]
           not ‹X ⊆ S› w
      show False by auto
    qed
    ultimately show ?thesis
      by (rule *)
  qed
  then have "open (f ` ⋃components U)"
    by (metis (no_types, lifting) imageE image_Union open_Union)
  then show ?thesis
    by force
qed


text‹No need for @{term S} to be connected. But the nonconstant condition is stronger.›
corollary open_mapping_thm2:
  assumes holf: "f holomorphic_on S"
      and S: "open S"
      and "open U" "U ⊆ S"
      and fnc: "⋀X. ⟦open X; X ⊆ S; X ≠ {}⟧ ⟹ ~ f constant_on X"
    shows "open (f ` U)"
proof -
  have "S = ⋃(components S)" by simp
  with ‹U ⊆ S› have "U = (⋃C ∈ components S. C ∩ U)" by auto
  then have "f ` U = (⋃C ∈ components S. f ` (C ∩ U))"
    using image_UN by fastforce
  moreover
  { fix C assume "C ∈ components S"
    with S ‹C ∈ components S› open_components in_components_connected
    have C: "open C" "connected C" by auto
    have "C ⊆ S"
      by (metis ‹C ∈ components S› in_components_maximal)
    have nf: "¬ f constant_on C"
      apply (rule fnc)
      using C ‹C ⊆ S› ‹C ∈ components S› in_components_nonempty by auto
    have "f holomorphic_on C"
      by (metis holf holomorphic_on_subset ‹C ⊆ S›)
    then have "open (f ` (C ∩ U))"
      apply (rule open_mapping_thm [OF _ C _ _ nf])
      apply (simp add: C ‹open U› open_Int, blast)
      done
  } ultimately show ?thesis
    by force
qed

corollary open_mapping_thm3:
  assumes holf: "f holomorphic_on S"
      and "open S" and injf: "inj_on f S"
    shows  "open (f ` S)"
apply (rule open_mapping_thm2 [OF holf])
using assms
apply (simp_all add:)
using injective_not_constant subset_inj_on by blast



subsection‹Maximum Modulus Principle›

text‹If @{term f} is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is
   properly within the domain of @{term f}.›

proposition maximum_modulus_principle:
  assumes holf: "f holomorphic_on S"
      and S: "open S" "connected S"
      and "open U" "U ⊆ S" "ξ ∈ U"
      and no: "⋀z. z ∈ U ⟹ norm(f z) ≤ norm(f ξ)"
    shows "f constant_on S"
proof (rule ccontr)
  assume "¬ f constant_on S"
  then have "open (f ` U)"
    using open_mapping_thm assms by blast
  moreover have "~ open (f ` U)"
  proof -
    have "∃t. cmod (f ξ - t) < e ∧ t ∉ f ` U" if "0 < e" for e
      apply (rule_tac x="if 0 < Re(f ξ) then f ξ + (e/2) else f ξ - (e/2)" in exI)
      using that
      apply (simp add: dist_norm)
      apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym)
      done
    then show ?thesis
      unfolding open_contains_ball by (metis ‹ξ ∈ U› contra_subsetD dist_norm imageI mem_ball)
  qed
  ultimately show False
    by blast
qed


proposition maximum_modulus_frontier:
  assumes holf: "f holomorphic_on (interior S)"
      and contf: "continuous_on (closure S) f"
      and bos: "bounded S"
      and leB: "⋀z. z ∈ frontier S ⟹ norm(f z) ≤ B"
      and "ξ ∈ S"
    shows "norm(f ξ) ≤ B"
proof -
  have "compact (closure S)" using bos
    by (simp add: bounded_closure compact_eq_bounded_closed)
  moreover have "continuous_on (closure S) (cmod ∘ f)"
    using contf continuous_on_compose continuous_on_norm_id by blast
  ultimately obtain z where zin: "z ∈ closure S" and z: "⋀y. y ∈ closure S ⟹ (cmod ∘ f) y ≤ (cmod ∘ f) z"
    using continuous_attains_sup [of "closure S" "norm o f"] ‹ξ ∈ S› by auto
  then consider "z ∈ frontier S" | "z ∈ interior S" using frontier_def by auto
  then have "norm(f z) ≤ B"
  proof cases
    case 1 then show ?thesis using leB by blast
  next
    case 2
    have zin: "z ∈ connected_component_set (interior S) z"
      by (simp add: 2)
    have "f constant_on (connected_component_set (interior S) z)"
      apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin])
      apply (metis connected_component_subset holf holomorphic_on_subset)
      apply (simp_all add: open_connected_component)
      by (metis closure_subset comp_eq_dest_lhs  interior_subset subsetCE z connected_component_in)
    then obtain c where c: "⋀w. w ∈ connected_component_set (interior S) z ⟹ f w = c"
      by (auto simp: constant_on_def)
    have "f ` closure(connected_component_set (interior S) z) ⊆ {c}"
      apply (rule image_closure_subset)
      apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset)
      using c
      apply auto
      done
    then have cc: "⋀w. w ∈ closure(connected_component_set (interior S) z) ⟹ f w = c" by blast
    have "frontier(connected_component_set (interior S) z) ≠ {}"
      apply (simp add: frontier_eq_empty)
      by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV)
    then obtain w where w: "w ∈ frontier(connected_component_set (interior S) z)"
       by auto
    then have "norm (f z) = norm (f w)"  by (simp add: "2" c cc frontier_def)
    also have "... ≤ B"
      apply (rule leB)
      using w
using frontier_interior_subset frontier_of_connected_component_subset by blast
    finally show ?thesis .
  qed
  then show ?thesis
    using z ‹ξ ∈ S› closure_subset by fastforce
qed

corollary maximum_real_frontier:
  assumes holf: "f holomorphic_on (interior S)"
      and contf: "continuous_on (closure S) f"
      and bos: "bounded S"
      and leB: "⋀z. z ∈ frontier S ⟹ Re(f z) ≤ B"
      and "ξ ∈ S"
    shows "Re(f ξ) ≤ B"
using maximum_modulus_frontier [of "exp o f" S "exp B"]
      Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms
by auto


subsection‹Factoring out a zero according to its order›

lemma holomorphic_factor_order_of_zero:
  assumes holf: "f holomorphic_on S"
      and os: "open S"
      and "ξ ∈ S" "0 < n"
      and dnz: "(deriv ^^ n) f ξ ≠ 0"
      and dfz: "⋀i. ⟦0 < i; i < n⟧ ⟹ (deriv ^^ i) f ξ = 0"
   obtains g r where "0 < r"
                "g holomorphic_on ball ξ r"
                "⋀w. w ∈ ball ξ r ⟹ f w - f ξ = (w - ξ)^n * g w"
                "⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
proof -
  obtain r where "r>0" and r: "ball ξ r ⊆ S" using assms by (blast elim!: openE)
  then have holfb: "f holomorphic_on ball ξ r"
    using holf holomorphic_on_subset by blast
  define g where "g w = suminf (λi. (deriv ^^ (i + n)) f ξ / (fact(i + n)) * (w - ξ)^i)" for w
  have sumsg: "(λi. (deriv ^^ (i + n)) f ξ / (fact(i + n)) * (w - ξ)^i) sums g w"
   and feq: "f w - f ξ = (w - ξ)^n * g w"
       if w: "w ∈ ball ξ r" for w
  proof -
    define powf where "powf = (λi. (deriv ^^ i) f ξ/(fact i) * (w - ξ)^i)"
    have sing: "{..<n} - {i. powf i = 0} = (if f ξ = 0 then {} else {0})"
      unfolding powf_def using ‹0 < n› dfz by (auto simp: dfz; metis funpow_0 not_gr0)
    have "powf sums f w"
      unfolding powf_def by (rule holomorphic_power_series [OF holfb w])
    moreover have "(∑i<n. powf i) = f ξ"
      apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric])
      apply (simp add:)
      apply (simp only: dfz sing)
      apply (simp add: powf_def)
      done
    ultimately have fsums: "(λi. powf (i+n)) sums (f w - f ξ)"
      using w sums_iff_shift' by metis
    then have *: "summable (λi. (w - ξ) ^ n * ((deriv ^^ (i + n)) f ξ * (w - ξ) ^ i / fact (i + n)))"
      unfolding powf_def using sums_summable
      by (auto simp: power_add mult_ac)
    have "summable (λi. (deriv ^^ (i + n)) f ξ * (w - ξ) ^ i / fact (i + n))"
    proof (cases "w=ξ")
      case False then show ?thesis
        using summable_mult [OF *, of "1 / (w - ξ) ^ n"] by (simp add:)
    next
      case True then show ?thesis
        by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"]
                 split: if_split_asm)
    qed
    then show sumsg: "(λi. (deriv ^^ (i + n)) f ξ / (fact(i + n)) * (w - ξ)^i) sums g w"
      by (simp add: summable_sums_iff g_def)
    show "f w - f ξ = (w - ξ)^n * g w"
      apply (rule sums_unique2)
      apply (rule fsums [unfolded powf_def])
      using sums_mult [OF sumsg, of "(w - ξ) ^ n"]
      by (auto simp: power_add mult_ac)
  qed
  then have holg: "g holomorphic_on ball ξ r"
    by (meson sumsg power_series_holomorphic)
  then have contg: "continuous_on (ball ξ r) g"
    by (blast intro: holomorphic_on_imp_continuous_on)
  have "g ξ ≠ 0"
    using dnz unfolding g_def
    by (subst suminf_finite [of "{0}"]) auto
  obtain d where "0 < d" and d: "⋀w. w ∈ ball ξ d ⟹ g w ≠ 0"
    apply (rule exE [OF continuous_on_avoid [OF contg _ ‹g ξ ≠ 0›]])
    using ‹0 < r›
    apply force
    by (metis ‹0 < r› less_trans mem_ball not_less_iff_gr_or_eq)
  show ?thesis
    apply (rule that [where g=g and r ="min r d"])
    using ‹0 < r› ‹0 < d› holg
    apply (auto simp: feq holomorphic_on_subset subset_ball d)
    done
qed


lemma holomorphic_factor_order_of_zero_strong:
  assumes holf: "f holomorphic_on S" "open S"  "ξ ∈ S" "0 < n"
      and "(deriv ^^ n) f ξ ≠ 0"
      and "⋀i. ⟦0 < i; i < n⟧ ⟹ (deriv ^^ i) f ξ = 0"
   obtains g r where "0 < r"
                "g holomorphic_on ball ξ r"
                "⋀w. w ∈ ball ξ r ⟹ f w - f ξ = ((w - ξ) * g w) ^ n"
                "⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
proof -
  obtain g r where "0 < r"
               and holg: "g holomorphic_on ball ξ r"
               and feq: "⋀w. w ∈ ball ξ r ⟹ f w - f ξ = (w - ξ)^n * g w"
               and gne: "⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
    by (auto intro: holomorphic_factor_order_of_zero [OF assms])
  have con: "continuous_on (ball ξ r) (λz. deriv g z / g z)"
    by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on)
  have cd: "⋀x. dist ξ x < r ⟹ (λz. deriv g z / g z) field_differentiable at x"
    apply (rule derivative_intros)+
    using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at)
    apply (metis Topology_Euclidean_Space.open_ball at_within_open holg holomorphic_on_def mem_ball)
    using gne mem_ball by blast
  obtain h where h: "⋀x. x ∈ ball ξ r ⟹ (h has_field_derivative deriv g x / g x) (at x)"
    apply (rule exE [OF holomorphic_convex_primitive [of "ball ξ r" "{}" "λz. deriv g z / g z"]])
    apply (auto simp: con cd)
    apply (metis open_ball at_within_open mem_ball)
    done
  then have "continuous_on (ball ξ r) h"
    by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open)
  then have con: "continuous_on (ball ξ r) (λx. exp (h x) / g x)"
    by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne)
  have 0: "dist ξ x < r ⟹ ((λx. exp (h x) / g x) has_field_derivative 0) (at x)" for x
    apply (rule h derivative_eq_intros | simp)+
    apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2])
    using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h)
    done
  obtain c where c: "⋀x. x ∈ ball ξ r ⟹ exp (h x) / g x = c"
    by (rule DERIV_zero_connected_constant [of "ball ξ r" "{}" "λx. exp(h x) / g x"]) (auto simp: con 0)
  have hol: "(λz. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball ξ r"
    apply (rule holomorphic_on_compose [unfolded o_def, where g = exp])
    apply (rule holomorphic_intros)+
    using h holomorphic_on_open apply blast
    apply (rule holomorphic_intros)+
    using ‹0 < n› apply (simp add:)
    apply (rule holomorphic_intros)+
    done
  show ?thesis
    apply (rule that [where g="λz. exp((Ln(inverse c) + h z)/n)" and r =r])
    using ‹0 < r› ‹0 < n›
    apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric])
    apply (rule hol)
    apply (simp add: Transcendental.exp_add gne)
    done
qed


lemma
  fixes k :: "'a::wellorder"
  assumes a_def: "a == LEAST x. P x" and P: "P k"
  shows def_LeastI: "P a" and def_Least_le: "a ≤ k"
unfolding a_def
by (rule LeastI Least_le; rule P)+

lemma holomorphic_factor_zero_nonconstant:
  assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
      and "ξ ∈ S" "f ξ = 0"
      and nonconst: "⋀c. ∃z ∈ S. f z ≠ c"
   obtains g r n
      where "0 < n"  "0 < r"  "ball ξ r ⊆ S"
            "g holomorphic_on ball ξ r"
            "⋀w. w ∈ ball ξ r ⟹ f w = (w - ξ)^n * g w"
            "⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
proof (cases "∀n>0. (deriv ^^ n) f ξ = 0")
  case True then show ?thesis
    using holomorphic_fun_eq_const_on_connected [OF holf S _ ‹ξ ∈ S›] nonconst by auto
next
  case False
  then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f ξ ≠ 0" by blast
  obtain r0 where "r0 > 0" "ball ξ r0 ⊆ S" using S openE ‹ξ ∈ S› by auto
  define n where "n ≡ LEAST n. (deriv ^^ n) f ξ ≠ 0"
  have n_ne: "(deriv ^^ n) f ξ ≠ 0"
    by (rule def_LeastI [OF n_def]) (rule n0)
  then have "0 < n" using ‹f ξ = 0›
    using funpow_0 by fastforce
  have n_min: "⋀k. k < n ⟹ (deriv ^^ k) f ξ = 0"
    using def_Least_le [OF n_def] not_le by blast
  then obtain g r1
    where  "0 < r1" "g holomorphic_on ball ξ r1"
           "⋀w. w ∈ ball ξ r1 ⟹ f w = (w - ξ) ^ n * g w"
           "⋀w. w ∈ ball ξ r1 ⟹ g w ≠ 0"
    by (auto intro: holomorphic_factor_order_of_zero [OF holf ‹open S› ‹ξ ∈ S› ‹n > 0› n_ne] simp: ‹f ξ = 0›)
  then show ?thesis
    apply (rule_tac g=g and r="min r0 r1" and n=n in that)
    using ‹0 < n› ‹0 < r0› ‹0 < r1› ‹ball ξ r0 ⊆ S›
    apply (auto simp: subset_ball intro: holomorphic_on_subset)
    done
qed


lemma holomorphic_lower_bound_difference:
  assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
      and "ξ ∈ S" and "φ ∈ S"
      and fne: "f φ ≠ f ξ"
   obtains k n r
      where "0 < k"  "0 < r"
            "ball ξ r ⊆ S"
            "⋀w. w ∈ ball ξ r ⟹ k * norm(w - ξ)^n ≤ norm(f w - f ξ)"
proof -
  define n where "n = (LEAST n. 0 < n ∧ (deriv ^^ n) f ξ ≠ 0)"
  obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f ξ ≠ 0"
    using fne holomorphic_fun_eq_const_on_connected [OF holf S] ‹ξ ∈ S› ‹φ ∈ S› by blast
  then have "0 < n" and n_ne: "(deriv ^^ n) f ξ ≠ 0"
    unfolding n_def by (metis (mono_tags, lifting) LeastI)+
  have n_min: "⋀k. ⟦0 < k; k < n⟧ ⟹ (deriv ^^ k) f ξ = 0"
    unfolding n_def by (blast dest: not_less_Least)
  then obtain g r
    where "0 < r" and holg: "g holomorphic_on ball ξ r"
      and fne: "⋀w. w ∈ ball ξ r ⟹ f w - f ξ = (w - ξ) ^ n * g w"
      and gnz: "⋀w. w ∈ ball ξ r ⟹ g w ≠ 0"
      by (auto intro: holomorphic_factor_order_of_zero  [OF holf ‹open S› ‹ξ ∈ S› ‹n > 0› n_ne])
  obtain e where "e>0" and e: "ball ξ e ⊆ S" using assms by (blast elim!: openE)
  then have holfb: "f holomorphic_on ball ξ e"
    using holf holomorphic_on_subset by blast
  define d where "d = (min e r) / 2"
  have "0 < d" using ‹0 < r› ‹0 < e› by (simp add: d_def)
  have "d < r"
    using ‹0 < r› by (auto simp: d_def)
  then have cbb: "cball ξ d ⊆ ball ξ r"
    by (auto simp: cball_subset_ball_iff)
  then have "g holomorphic_on cball ξ d"
    by (rule holomorphic_on_subset [OF holg])
  then have "closed (g ` cball ξ d)"
    by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on)
  moreover have "g ` cball ξ d ≠ {}"
    using ‹0 < d› by auto
  ultimately obtain x where x: "x ∈ g ` cball ξ d" and "⋀y. y ∈ g ` cball ξ d ⟹ dist 0 x ≤ dist 0 y"
    by (rule distance_attains_inf) blast
  then have leg: "⋀w. w ∈ cball ξ d ⟹ norm x ≤ norm (g w)"
    by auto
  have "ball ξ d ⊆ cball ξ d" by auto
  also have "... ⊆ ball ξ e" using ‹0 < d› d_def by auto
  also have "... ⊆ S" by (rule e)
  finally have dS: "ball ξ d ⊆ S" .
  moreover have "x ≠ 0" using gnz x ‹d < r› by auto
  ultimately show ?thesis
    apply (rule_tac k="norm x" and n=n and r=d in that)
    using ‹d < r› leg
    apply (auto simp: ‹0 < d› fne norm_mult norm_power algebra_simps mult_right_mono)
    done
qed

lemma
  assumes holf: "f holomorphic_on (S - {ξ})" and ξ: "ξ ∈ interior S"
    shows holomorphic_on_extend_lim:
          "(∃g. g holomorphic_on S ∧ (∀z ∈ S - {ξ}. g z = f z)) ⟷
           ((λz. (z - ξ) * f z) ⤏ 0) (at ξ)"
          (is "?P = ?Q")
     and holomorphic_on_extend_bounded:
          "(∃g. g holomorphic_on S ∧ (∀z ∈ S - {ξ}. g z = f z)) ⟷
           (∃B. eventually (λz. norm(f z) ≤ B) (at ξ))"
          (is "?P = ?R")
proof -
  obtain δ where "0 < δ" and δ: "ball ξ δ ⊆ S"
    using ξ mem_interior by blast
  have "?R" if holg: "g holomorphic_on S" and gf: "⋀z. z ∈ S - {ξ} ⟹ g z = f z" for g
  proof -
    have *: "∀F z in at ξ. dist (g z) (g ξ) < 1 ⟶ cmod (f z) ≤ cmod (g ξ) + 1"
      apply (simp add: eventually_at)
      apply (rule_tac x="δ" in exI)
      using δ ‹0 < δ›
      apply (clarsimp simp:)
      apply (drule_tac c=x in subsetD)
      apply (simp add: dist_commute)
      by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD)
    have "continuous_on (interior S) g"
      by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset)
    then have "⋀x. x ∈ interior S ⟹ (g ⤏ g x) (at x)"
      using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast
    then have "(g ⤏ g ξ) (at ξ)"
      by (simp add: ξ)
    then show ?thesis
      apply (rule_tac x="norm(g ξ) + 1" in exI)
      apply (rule eventually_mp [OF * tendstoD [where e=1]], auto)
      done
  qed
  moreover have "?Q" if "∀F z in at ξ. cmod (f z) ≤ B" for B
    by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero)
  moreover have "?P" if "(λz. (z - ξ) * f z) ─ξ→ 0"
  proof -
    define h where [abs_def]: "h z = (z - ξ)^2 * f z" for z
    have h0: "(h has_field_derivative 0) (at ξ)"
      apply (simp add: h_def Derivative.DERIV_within_iff)
      apply (rule Lim_transform_within [OF that, of 1])
      apply (auto simp: divide_simps power2_eq_square)
      done
    have holh: "h holomorphic_on S"
    proof (simp add: holomorphic_on_def, clarify)
      fix z assume "z ∈ S"
      show "h field_differentiable at z within S"
      proof (cases "z = ξ")
        case True then show ?thesis
          using field_differentiable_at_within field_differentiable_def h0 by blast
      next
        case False
        then have "f field_differentiable at z within S"
          using holomorphic_onD [OF holf, of z] ‹z ∈ S›
          unfolding field_differentiable_def DERIV_within_iff
          by (force intro: exI [where x="dist ξ z"] elim: Lim_transform_within_set [unfolded eventually_at])
        then show ?thesis
          by (simp add: h_def power2_eq_square derivative_intros)
      qed
    qed
    define g where [abs_def]: "g z = (if z = ξ then deriv h ξ else (h z - h ξ) / (z - ξ))" for z
    have holg: "g holomorphic_on S"
      unfolding g_def by (rule pole_lemma [OF holh ξ])
    show ?thesis
      apply (rule_tac x="λz. if z = ξ then deriv g ξ else (g z - g ξ)/(z - ξ)" in exI)
      apply (rule conjI)
      apply (rule pole_lemma [OF holg ξ])
      apply (auto simp: g_def power2_eq_square divide_simps)
      using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square)
      done
  qed
  ultimately show "?P = ?Q" and "?P = ?R"
    by meson+
qed


proposition pole_at_infinity:
  assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) ⤏ l) at_infinity"
  obtains a n where "⋀z. f z = (∑i≤n. a i * z^i)"
proof (cases "l = 0")
  case False
  with tendsto_inverse [OF lim] show ?thesis
    apply (rule_tac a="(λn. inverse l)" and n=0 in that)
    apply (simp add: Liouville_weak [OF holf, of "inverse l"])
    done
next
  case True
  then have [simp]: "l = 0" .
  show ?thesis
  proof (cases "∃r. 0 < r ∧ (∀z ∈ ball 0 r - {0}. f(inverse z) ≠ 0)")
    case True
      then obtain r where "0 < r" and r: "⋀z. z ∈ ball 0 r - {0} ⟹ f(inverse z) ≠ 0"
             by auto
      have 1: "inverse ∘ f ∘ inverse holomorphic_on ball 0 r - {0}"
        by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+
      have 2: "0 ∈ interior (ball 0 r)"
        using ‹0 < r› by simp
      have "∃B. 0<B ∧ eventually (λz. cmod ((inverse ∘ f ∘ inverse) z) ≤ B) (at 0)"
        apply (rule exI [where x=1])
        apply (simp add:)
        using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one]
        apply (rule eventually_mono)
        apply (simp add: dist_norm)
        done
      with holomorphic_on_extend_bounded [OF 1 2]
      obtain g where holg: "g holomorphic_on ball 0 r"
                 and geq: "⋀z. z ∈ ball 0 r - {0} ⟹ g z = (inverse ∘ f ∘ inverse) z"
        by meson
      have ifi0: "(inverse ∘ f ∘ inverse) ─0→ 0"
        using ‹l = 0› lim lim_at_infinity_0 by blast
      have g2g0: "g ─0→ g 0"
        using ‹0 < r› centre_in_ball continuous_at continuous_on_eq_continuous_at holg
        by (blast intro: holomorphic_on_imp_continuous_on)
      have g2g1: "g ─0→ 0"
        apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]])
        using ‹0 < r› by (auto simp: geq)
      have [simp]: "g 0 = 0"
        by (rule tendsto_unique [OF _ g2g0 g2g1]) simp
      have "ball 0 r - {0::complex} ≠ {}"
        using ‹0 < r›
        apply (clarsimp simp: ball_def dist_norm)
        apply (drule_tac c="of_real r/2" in subsetD, auto)
        done
      then obtain w::complex where "w ≠ 0" and w: "norm w < r" by force
      then have "g w ≠ 0" by (simp add: geq r)
      obtain B n e where "0 < B" "0 < e" "e ≤ r"
                     and leg: "⋀w. norm w < e ⟹ B * cmod w ^ n ≤ cmod (g w)"
        apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w])
        using ‹0 < r› w ‹g w ≠ 0› by (auto simp: ball_subset_ball_iff)
      have "cmod (f z) ≤ cmod z ^ n / B" if "2/e ≤ cmod z" for z
      proof -
        have ize: "inverse z ∈ ball 0 e - {0}" using that ‹0 < e›
          by (auto simp: norm_divide divide_simps algebra_simps)
        then have [simp]: "z ≠ 0" and izr: "inverse z ∈ ball 0 r - {0}" using  ‹e ≤ r›
          by auto
        then have [simp]: "f z ≠ 0"
          using r [of "inverse z"] by simp
        have [simp]: "f z = inverse (g (inverse z))"
          using izr geq [of "inverse z"] by simp
        show ?thesis using ize leg [of "inverse z"]  ‹0 < B›  ‹0 < e›
          by (simp add: divide_simps norm_divide algebra_simps)
      qed
      then show ?thesis
        apply (rule_tac a = "λk. (deriv ^^ k) f 0 / (fact k)" and n=n in that)
        apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf])
        apply (simp add:)
        done
  next
    case False
    then have fi0: "⋀r. r > 0 ⟹ ∃z∈ball 0 r - {0}. f (inverse z) = 0"
      by simp
    have fz0: "f z = 0" if "0 < r" and lt1: "⋀x. x ≠ 0 ⟹ cmod x < r ⟹ inverse (cmod (f (inverse x))) < 1"
              for z r
    proof -
      have f0: "(f ⤏ 0) at_infinity"
      proof -
        have DIM_complex[intro]: "2 ≤ DIM(complex)"  ‹should not be necessary!›
          by simp
        have "continuous_on (inverse ` (ball 0 r - {0})) f"
          using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast
        then have "connected ((f ∘ inverse) ` (ball 0 r - {0}))"
          apply (intro connected_continuous_image continuous_intros)
          apply (force intro: connected_punctured_ball)+
          done
        then have "⟦w ≠ 0; cmod w < r⟧ ⟹ f (inverse w) = 0" for w
          apply (rule disjE [OF connected_closedD [where A = "{0}" and B = "- ball 0 1"]], auto)
          apply (metis (mono_tags, hide_lams) not_less_iff_gr_or_eq one_less_inverse lt1 zero_less_norm_iff)
          using False ‹0 < r› apply fastforce
          by (metis (no_types, hide_lams) Compl_iff IntI comp_apply empty_iff image_eqI insert_Diff_single insert_iff mem_ball_0 not_less_iff_gr_or_eq one_less_inverse that(2) zero_less_norm_iff)
        then show ?thesis
          apply (simp add: lim_at_infinity_0)
          apply (rule Lim_eventually)
          apply (simp add: eventually_at)
          apply (rule_tac x=r in exI)
          apply (simp add: ‹0 < r› dist_norm)
          done
      qed
      obtain w where "w ∈ ball 0 r - {0}" and "f (inverse w) = 0"
        using False ‹0 < r› by blast
      then show ?thesis
        by (auto simp: f0 Liouville_weak [OF holf, of 0])
    qed
    show ?thesis
      apply (rule that [of "λn. 0" 0])
      using lim [unfolded lim_at_infinity_0]
      apply (simp add: Lim_at dist_norm norm_inverse)
      apply (drule_tac x=1 in spec)
      using fz0 apply auto
      done
    qed
qed


subsection‹Entire proper functions are precisely the non-trivial polynomials›

proposition proper_map_polyfun:
    fixes c :: "nat ⇒ 'a::{real_normed_div_algebra,heine_borel}"
  assumes "closed S" and "compact K" and c: "c i ≠ 0" "1 ≤ i" "i ≤ n"
    shows "compact (S ∩ {z. (∑i≤n. c i * z^i) ∈ K})"
proof -
  obtain B where "B > 0" and B: "⋀x. x ∈ K ⟹ norm x ≤ B"
    by (metis compact_imp_bounded ‹compact K› bounded_pos)
  have *: "norm x ≤ b"
            if "⋀x. b ≤ norm x ⟹ B + 1 ≤ norm (∑i≤n. c i * x ^ i)"
               "(∑i≤n. c i * x ^ i) ∈ K"  for b x
  proof -
    have "norm (∑i≤n. c i * x ^ i) ≤ B"
      using B that by blast
    moreover have "¬ B + 1 ≤ B"
      by simp
    ultimately show "norm x ≤ b"
      using that by (metis (no_types) less_eq_real_def not_less order_trans)
  qed
  have "bounded {z. (∑i≤n. c i * z ^ i) ∈ K}"
    using polyfun_extremal [where c=c and B="B+1", OF c]
    by (auto simp: bounded_pos eventually_at_infinity_pos *)
  moreover have "closed {z. (∑i≤n. c i * z ^ i) ∈ K}"
    apply (intro allI continuous_closed_preimage_univ continuous_intros)
    using ‹compact K› compact_eq_bounded_closed by blast
  ultimately show ?thesis
    using closed_Int_compact [OF ‹closed S›] compact_eq_bounded_closed by blast
qed

corollary proper_map_polyfun_univ:
    fixes c :: "nat ⇒ 'a::{real_normed_div_algebra,heine_borel}"
  assumes "compact K" "c i ≠ 0" "1 ≤ i" "i ≤ n"
    shows "compact ({z. (∑i≤n. c i * z^i) ∈ K})"
using proper_map_polyfun [of UNIV K c i n] assms by simp


proposition proper_map_polyfun_eq:
  assumes "f holomorphic_on UNIV"
    shows "(∀k. compact k ⟶ compact {z. f z ∈ k}) ⟷
           (∃c n. 0 < n ∧ (c n ≠ 0) ∧ f = (λz. ∑i≤n. c i * z^i))"
          (is "?lhs = ?rhs")
proof
  assume compf [rule_format]: ?lhs
  have 2: "∃k. 0 < k ∧ a k ≠ 0 ∧ f = (λz. ∑i ≤ k. a i * z ^ i)"
        if "⋀z. f z = (∑i≤n. a i * z ^ i)" for a n
  proof (cases "∀i≤n. 0<i ⟶ a i = 0")
    case True
    then have [simp]: "⋀z. f z = a 0"
      by (simp add: that sum_atMost_shift)
    have False using compf [of "{a 0}"] by simp
    then show ?thesis ..
  next
    case False
    then obtain k where k: "0 < k" "k≤n" "a k ≠ 0" by force
    define m where "m = (GREATEST k. k≤n ∧ a k ≠ 0)"
    have m: "m≤n ∧ a m ≠ 0"
      unfolding m_def
      apply (rule GreatestI_nat [where b = n])
      using k apply auto
      done
    have [simp]: "a i = 0" if "m < i" "i ≤ n" for i
      using Greatest_le_nat [where b = "n" and P = "λk. k≤n ∧ a k ≠ 0"]
      using m_def not_le that by auto
    have "k ≤ m"
      unfolding m_def
      apply (rule Greatest_le_nat [where b = "n"])
      using k apply auto
      done
    with k m show ?thesis
      by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right)
  qed
  have "((inverse ∘ f) ⤏ 0) at_infinity"
  proof (rule Lim_at_infinityI)
    fix e::real assume "0 < e"
    with compf [of "cball 0 (inverse e)"]
    show "∃B. ∀x. B ≤ cmod x ⟶ dist ((inverse ∘ f) x) 0 ≤ e"
      apply (simp add:)
      apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse)
      apply (rule_tac x="b+1" in exI)
      apply (metis inverse_inverse_eq less_add_same_cancel2 less_imp_inverse_less add.commute not_le not_less_iff_gr_or_eq order_trans zero_less_one)
      done
  qed
  then show ?rhs
    apply (rule pole_at_infinity [OF assms])
    using 2 apply blast
    done
next
  assume ?rhs
  then obtain c n where "0 < n" "c n ≠ 0" "f = (λz. ∑i≤n. c i * z ^ i)" by blast
  then have "compact {z. f z ∈ k}" if "compact k" for k
    by (auto intro: proper_map_polyfun_univ [OF that])
  then show ?lhs by blast
qed


subsection‹Relating invertibility and nonvanishing of derivative›

proposition has_complex_derivative_locally_injective:
  assumes holf: "f holomorphic_on S"
      and S: "ξ ∈ S" "open S"
      and dnz: "deriv f ξ ≠ 0"
  obtains r where "r > 0" "ball ξ r ⊆ S" "inj_on f (ball ξ r)"
proof -
  have *: "∃d>0. ∀x. dist ξ x < d ⟶ onorm (λv. deriv f x * v - deriv f ξ * v) < e" if "e > 0" for e
  proof -
    have contdf: "continuous_on S (deriv f)"
      by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on ‹open S›)
    obtain δ where "δ>0" and δ: "⋀x. ⟦x ∈ S; dist x ξ ≤ δ⟧ ⟹ cmod (deriv f x - deriv f ξ) ≤ e/2"
      using continuous_onE [OF contdf ‹ξ ∈ S›, of "e/2"] ‹0 < e›
      by (metis dist_complex_def half_gt_zero less_imp_le)
    obtain ε where "ε>0" "ball ξ ε ⊆ S"
      by (metis openE [OF ‹open S› ‹ξ ∈ S›])
    with ‹δ>0› have "∃δ>0. ∀x. dist ξ x < δ ⟶ onorm (λv. deriv f x * v - deriv f ξ * v) ≤ e/2"
      apply (rule_tac x="min δ ε" in exI)
      apply (intro conjI allI impI Operator_Norm.onorm_le)
      apply (simp add:)
      apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult)
      apply (rule mult_right_mono [OF δ])
      apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono δ)
      done
    with ‹e>0› show ?thesis by force
  qed
  have "inj (op * (deriv f ξ))"
    using dnz by simp
  then obtain g' where g': "linear g'" "g' ∘ op * (deriv f ξ) = id"
    using linear_injective_left_inverse [of "op * (deriv f ξ)"]
    by (auto simp: linear_times)
  show ?thesis
    apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "λz h. deriv f z * h" and g' = g'])
    using g' *
    apply (simp_all add: linear_conv_bounded_linear that)
    using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf
        holomorphic_on_imp_differentiable_at ‹open S› apply blast
    done
qed


proposition has_complex_derivative_locally_invertible:
  assumes holf: "f holomorphic_on S"
      and S: "ξ ∈ S" "open S"
      and dnz: "deriv f ξ ≠ 0"
  obtains r where "r > 0" "ball ξ r ⊆ S" "open (f `  (ball ξ r))" "inj_on f (ball ξ r)"
proof -
  obtain r where "r > 0" "ball ξ r ⊆ S" "inj_on f (ball ξ r)"
    by (blast intro: that has_complex_derivative_locally_injective [OF assms])
  then have ξ: "ξ ∈ ball ξ r" by simp
  then have nc: "~ f constant_on ball ξ r"
    using ‹inj_on f (ball ξ r)› injective_not_constant by fastforce
  have holf': "f holomorphic_on ball ξ r"
    using ‹ball ξ r ⊆ S› holf holomorphic_on_subset by blast
  have "open (f ` ball ξ r)"
    apply (rule open_mapping_thm [OF holf'])
    using nc apply auto
    done
  then show ?thesis
    using ‹0 < r› ‹ball ξ r ⊆ S› ‹inj_on f (ball ξ r)› that  by blast
qed


proposition holomorphic_injective_imp_regular:
  assumes holf: "f holomorphic_on S"
      and "open S" and injf: "inj_on f S"
      and "ξ ∈ S"
    shows "deriv f ξ ≠ 0"
proof -
  obtain r where "r>0" and r: "ball ξ r ⊆ S" using assms by (blast elim!: openE)
  have holf': "f holomorphic_on ball ξ r"
    using ‹ball ξ r ⊆ S› holf holomorphic_on_subset by blast
  show ?thesis
  proof (cases "∀n>0. (deriv ^^ n) f ξ = 0")
    case True
    have fcon: "f w = f ξ" if "w ∈ ball ξ r" for w
      apply (rule holomorphic_fun_eq_const_on_connected [OF holf'])
      using True ‹0 < r› that by auto
    have False
      using fcon [of "ξ + r/2"] ‹0 < r› r injf unfolding inj_on_def
      by (metis ‹ξ ∈ S› contra_subsetD dist_commute fcon mem_ball perfect_choose_dist)
    then show ?thesis ..
  next
    case False
    then obtain n0 where n0: "n0 > 0 ∧ (deriv ^^ n0) f ξ ≠ 0" by blast
    define n where [abs_def]: "n = (LEAST n. n > 0 ∧ (deriv ^^ n) f ξ ≠ 0)"
    have n_ne: "n > 0" "(deriv ^^ n) f ξ ≠ 0"
      using def_LeastI [OF n_def n0] by auto
    have n_min: "⋀k. 0 < k ⟹ k < n ⟹ (deriv ^^ k) f ξ = 0"
      using def_Least_le [OF n_def] not_le by auto
    obtain g δ where "0 < δ"
             and holg: "g holomorphic_on ball ξ δ"
             and fd: "⋀w. w ∈ ball ξ δ ⟹ f w - f ξ = ((w - ξ) * g w) ^ n"
             and gnz: "⋀w. w ∈ ball ξ δ ⟹ g w ≠ 0"
      apply (rule holomorphic_factor_order_of_zero_strong [OF holf ‹open S› ‹ξ ∈ S› n_ne])
      apply (blast intro: n_min)+
      done
    show ?thesis
    proof (cases "n=1")
      case True
      with n_ne show ?thesis by auto
    next
      case False
      have holgw: "(λw. (w - ξ) * g w) holomorphic_on ball ξ (min r δ)"
        apply (rule holomorphic_intros)+
        using holg by (simp add: holomorphic_on_subset subset_ball)
      have gd: "⋀w. dist ξ w < δ ⟹ (g has_field_derivative deriv g w) (at w)"
        using holg
        by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH)
      have *: "⋀w. w ∈ ball ξ (min r δ)
            ⟹ ((λw. (w - ξ) * g w) has_field_derivative ((w - ξ) * deriv g w + g w))
                (at w)"
        by (rule gd derivative_eq_intros | simp)+
      have [simp]: "deriv (λw. (w - ξ) * g w) ξ ≠ 0"
        using * [of ξ] ‹0 < δ› ‹0 < r› by (simp add: DERIV_imp_deriv gnz)
      obtain T where "ξ ∈ T" "open T" and Tsb: "T ⊆ ball ξ (min r δ)" and oimT: "open ((λw. (w - ξ) * g w) ` T)"
        apply (rule has_complex_derivative_locally_invertible [OF holgw, of ξ])
        using ‹0 < r› ‹0 < δ›
        apply (simp_all add:)
        by (meson Topology_Euclidean_Space.open_ball centre_in_ball)
      define U where "U = (λw. (w - ξ) * g w) ` T"
      have "open U" by (metis oimT U_def)
      have "0 ∈ U"
        apply (auto simp: U_def)
        apply (rule image_eqI [where x = ξ])
        apply (auto simp: ‹ξ ∈ T›)
        done
      then obtain ε where "ε>0" and ε: "cball 0 ε ⊆ U"
        using ‹open U› open_contains_cball by blast
      then have "ε * exp(2 * of_real pi * 𝗂 * (0/n)) ∈ cball 0 ε"
                "ε * exp(2 * of_real pi * 𝗂 * (1/n)) ∈ cball 0 ε"
        by (auto simp: norm_mult)
      with ε have "ε * exp(2 * of_real pi * 𝗂 * (0/n)) ∈ U"
                  "ε * exp(2 * of_real pi * 𝗂 * (1/n)) ∈ U" by blast+
      then obtain y0 y1 where "y0 ∈ T" and y0: "(y0 - ξ) * g y0 = ε * exp(2 * of_real pi * 𝗂 * (0/n))"
                          and "y1 ∈ T" and y1: "(y1 - ξ) * g y1 = ε * exp(2 * of_real pi * 𝗂 * (1/n))"
        by (auto simp: U_def)
      then have "y0 ∈ ball ξ δ" "y1 ∈ ball ξ δ" using Tsb by auto
      moreover have "y0 ≠ y1"
        using y0 y1 ‹ε > 0› complex_root_unity_eq_1 [of n 1] ‹n > 0› False by auto
      moreover have "T ⊆ S"
        by (meson Tsb min.cobounded1 order_trans r subset_ball)
      ultimately have False
        using inj_onD [OF injf, of y0 y1] ‹y0 ∈ T› ‹y1 ∈ T›
        using fd [of y0] fd [of y1] complex_root_unity [of n 1] n_ne
        apply (simp add: y0 y1 power_mult_distrib)
        apply (force simp: algebra_simps)
        done
      then show ?thesis ..
    qed
  qed
qed


text‹Hence a nice clean inverse function theorem›

proposition holomorphic_has_inverse:
  assumes holf: "f holomorphic_on S"
      and "open S" and injf: "inj_on f S"
  obtains g where "g holomorphic_on (f ` S)"
                  "⋀z. z ∈ S ⟹ deriv f z * deriv g (f z) = 1"
                  "⋀z. z ∈ S ⟹ g(f z) = z"
proof -
  have ofs: "open (f ` S)"
    by (rule open_mapping_thm3 [OF assms])
  have contf: "continuous_on S f"
    by (simp add: holf holomorphic_on_imp_continuous_on)
  have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z ∈ S" for z
  proof -
    have 1: "(f has_field_derivative deriv f z) (at z)"
      using DERIV_deriv_iff_field_differentiable ‹z ∈ S› ‹open S› holf holomorphic_on_imp_differentiable_at
      by blast
    have 2: "deriv f z ≠ 0"
      using ‹z ∈ S› ‹open S› holf holomorphic_injective_imp_regular injf by blast
    show ?thesis
      apply (rule has_complex_derivative_inverse_strong [OF 1 2 ‹open S› ‹z ∈ S›])
       apply (simp add: holf holomorphic_on_imp_continuous_on)
      by (simp add: injf the_inv_into_f_f)
  qed
  show ?thesis
    proof
      show "the_inv_into S f holomorphic_on f ` S"
        by (simp add: holomorphic_on_open ofs) (blast intro: *)
    next
      fix z assume "z ∈ S"
      have "deriv f z ≠ 0"
        using ‹z ∈ S› ‹open S› holf holomorphic_injective_imp_regular injf by blast
      then show "deriv f z * deriv (the_inv_into S f) (f z) = 1"
        using * [OF ‹z ∈ S›]  by (simp add: DERIV_imp_deriv)
    next
      fix z assume "z ∈ S"
      show "the_inv_into S f (f z) = z"
        by (simp add: ‹z ∈ S› injf the_inv_into_f_f)
  qed
qed


subsection‹The Schwarz Lemma›

lemma Schwarz1:
  assumes holf: "f holomorphic_on S"
      and contf: "continuous_on (closure S) f"
      and S: "open S" "connected S"
      and boS: "bounded S"
      and "S ≠ {}"
  obtains w where "w ∈ frontier S"
                  "⋀z. z ∈ closure S ⟹ norm (f z) ≤ norm (f w)"
proof -
  have connf: "continuous_on (closure S) (norm o f)"
    using contf continuous_on_compose continuous_on_norm_id by blast
  have coc: "compact (closure S)"
    by (simp add: ‹bounded S› bounded_closure compact_eq_bounded_closed)
  then obtain x where x: "x ∈ closure S" and xmax: "⋀z. z ∈ closure S ⟹ norm(f z) ≤ norm(f x)"
    apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]])
    using ‹S ≠ {}› apply auto
    done
  then show ?thesis
  proof (cases "x ∈ frontier S")
    case True
    then show ?thesis using that xmax by blast
  next
    case False
    then have "x ∈ S"
      using ‹open S› frontier_def interior_eq x by auto
    then have "f constant_on S"
      apply (rule maximum_modulus_principle [OF holf S ‹open S› order_refl])
      using closure_subset apply (blast intro: xmax)
      done
    then have "f constant_on (closure S)"
      by (rule constant_on_closureI [OF _ contf])
    then obtain c where c: "⋀x. x ∈ closure S ⟹ f x = c"
      by (meson constant_on_def)
    obtain w where "w ∈ frontier S"
      by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV)
    then show ?thesis
      by (simp add: c frontier_def that)
  qed
qed

lemma Schwarz2:
 "⟦f holomorphic_on ball 0 r;
    0 < s; ball w s ⊆ ball 0 r;
    ⋀z. norm (w-z) < s ⟹ norm(f z) ≤ norm(f w)⟧
    ⟹ f constant_on ball 0 r"
by (rule maximum_modulus_principle [where U = "ball w s" and ξ = w]) (simp_all add: dist_norm)

lemma Schwarz3:
  assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0"
  obtains h where "h holomorphic_on (ball 0 r)" and "⋀z. norm z < r ⟹ f z = z * (h z)" and "deriv f 0 = h 0"
proof -
  define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z
  have d0: "deriv f 0 = h 0"
    by (simp add: h_def)
  moreover have "h holomorphic_on (ball 0 r)"
    by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def)
  moreover have "norm z < r ⟹ f z = z * h z" for z
    by (simp add: h_def)
  ultimately show ?thesis
    using that by blast
qed


proposition Schwarz_Lemma:
  assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0"
      and no: "⋀z. norm z < 1 ⟹ norm (f z) < 1"
      and ξ: "norm ξ < 1"
    shows "norm (f ξ) ≤ norm ξ" and "norm(deriv f 0) ≤ 1"
      and "((∃z. norm z < 1 ∧ z ≠ 0 ∧ norm(f z) = norm z) ∨ norm(deriv f 0) = 1)
           ⟹ ∃α. (∀z. norm z < 1 ⟶ f z = α * z) ∧ norm α = 1" (is "?P ⟹ ?Q")
proof -
  obtain h where holh: "h holomorphic_on (ball 0 1)"
             and fz_eq: "⋀z. norm z < 1 ⟹ f z = z * (h z)" and df0: "deriv f 0 = h 0"
    by (rule Schwarz3 [OF holf]) auto
  have noh_le: "norm (h z) ≤ 1" if z: "norm z < 1" for z
  proof -
    have "norm (h z) < a" if a: "1 < a" for a
    proof -
      have "max (inverse a) (norm z) < 1"
        using z a by (simp_all add: inverse_less_1_iff)
      then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1"
        using Rats_dense_in_real by blast
      then have nzr: "norm z < r" and ira: "inverse r < a"
        using z a less_imp_inverse_less by force+
      then have "0 < r"
        by (meson norm_not_less_zero not_le order.strict_trans2)
      have holh': "h holomorphic_on ball 0 r"
        by (meson holh ‹r < 1› holomorphic_on_subset less_eq_real_def subset_ball)
      have conth': "continuous_on (cball 0 r) h"
        by (meson ‹r < 1› dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI)
      obtain w where w: "norm w = r" and lenw: "⋀z. norm z < r ⟹ norm(h z) ≤ norm(h w)"
        apply (rule Schwarz1 [OF holh']) using conth' ‹0 < r› by auto
      have "h w = f w / w" using fz_eq ‹r < 1› nzr w by auto
      then have "cmod (h z) < inverse r"
        by (metis ‹0 < r› ‹r < 1› divide_strict_right_mono inverse_eq_divide
                  le_less_trans lenw no norm_divide nzr w)
      then show ?thesis using ira by linarith
    qed
    then show "norm (h z) ≤ 1"
      using not_le by blast
  qed
  show "cmod (f ξ) ≤ cmod ξ"
  proof (cases "ξ = 0")
    case True then show ?thesis by auto
  next
    case False
    then show ?thesis
      by (simp add: noh_le fz_eq ξ mult_left_le norm_mult)
  qed
  show no_df0: "norm(deriv f 0) ≤ 1"
    by (simp add: ‹⋀z. cmod z < 1 ⟹ cmod (h z) ≤ 1› df0)
  show "?Q" if "?P"
    using that
  proof
    assume "∃z. cmod z < 1 ∧ z ≠ 0 ∧ cmod (f z) = cmod z"
    then obtain γ where γ: "cmod γ < 1" "γ ≠ 0" "cmod (f γ) = cmod γ" by blast
    then have [simp]: "norm (h γ) = 1"
      by (simp add: fz_eq norm_mult)
    have "ball γ (1 - cmod γ) ⊆ ball 0 1"
      by (simp add: ball_subset_ball_iff)
    moreover have "⋀z. cmod (γ - z) < 1 - cmod γ ⟹ cmod (h z) ≤ cmod (h γ)"
      apply (simp add: algebra_simps)
      by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4)
    ultimately obtain c where c: "⋀z. norm z < 1 ⟹ h z = c"
      using Schwarz2 [OF holh, of "1 - norm γ" γ, unfolded constant_on_def] γ by auto
    then have "norm c = 1"
      using γ by force
    with c show ?thesis
      using fz_eq by auto
  next
    assume [simp]: "cmod (deriv f 0) = 1"
    then obtain c where c: "⋀z. norm z < 1 ⟹ h z = c"
      using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le
      by auto
    moreover have "norm c = 1"  using df0 c by auto
    ultimately show ?thesis
      using fz_eq by auto
  qed
qed

subsection‹The Schwarz reflection principle›

lemma hol_pal_lem0:
  assumes "d ∙ a ≤ k" "k ≤ d ∙ b"
  obtains c where
     "c ∈ closed_segment a b" "d ∙ c = k"
     "⋀z. z ∈ closed_segment a c ⟹ d ∙ z ≤ k"
     "⋀z. z ∈ closed_segment c b ⟹ k ≤ d ∙ z"
proof -
  obtain c where cin: "c ∈ closed_segment a b" and keq: "k = d ∙ c"
    using connected_ivt_hyperplane [of "closed_segment a b" a b d k]
    by (auto simp: assms)
  have "closed_segment a c ⊆ {z. d ∙ z ≤ k}"  "closed_segment c b ⊆ {z. k ≤ d ∙ z}"
    unfolding segment_convex_hull using assms keq
    by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal)
  then show ?thesis using cin that by fastforce
qed

lemma hol_pal_lem1:
  assumes "convex S" "open S"
      and abc: "a ∈ S" "b ∈ S" "c ∈ S"
          "d ≠ 0" and lek: "d ∙ a ≤ k" "d ∙ b ≤ k" "d ∙ c ≤ k"
      and holf1: "f holomorphic_on {z. z ∈ S ∧ d ∙ z < k}"
      and contf: "continuous_on S f"
    shows "contour_integral (linepath a b) f +
           contour_integral (linepath b c) f +
           contour_integral (linepath c a) f = 0"
proof -
  have "interior (convex hull {a, b, c}) ⊆ interior(S ∩ {x. d ∙ x ≤ k})"
    apply (rule interior_mono)
    apply (rule hull_minimal)
     apply (simp add: abc lek)
    apply (rule convex_Int [OF ‹convex S› convex_halfspace_le])
    done
  also have "... ⊆ {z ∈ S. d ∙ z < k}"
    by (force simp: interior_open [OF ‹open S›] ‹d ≠ 0›)
  finally have *: "interior (convex hull {a, b, c}) ⊆ {z ∈ S. d ∙ z < k}" .
  have "continuous_on (convex hull {a,b,c}) f"
    using ‹convex S› contf abc continuous_on_subset subset_hull
    by fastforce
  moreover have "f holomorphic_on interior (convex hull {a,b,c})"
    by (rule holomorphic_on_subset [OF holf1 *])
  ultimately show ?thesis
    using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3
      by blast
qed

lemma hol_pal_lem2:
  assumes S: "convex S" "open S"
      and abc: "a ∈ S" "b ∈ S" "c ∈ S"
      and "d ≠ 0" and lek: "d ∙ a ≤ k" "d ∙ b ≤ k"
      and holf1: "f holomorphic_on {z. z ∈ S ∧ d ∙ z < k}"
      and holf2: "f holomorphic_on {z. z ∈ S ∧ k < d ∙ z}"
      and contf: "continuous_on S f"
    shows "contour_integral (linepath a b) f +
           contour_integral (linepath b c) f +
           contour_integral (linepath c a) f = 0"
proof (cases "d ∙ c ≤ k")
  case True show ?thesis
    by (rule hol_pal_lem1 [OF S abc ‹d ≠ 0› lek True holf1 contf])
next
  case False
  then have "d ∙ c > k" by force
  obtain a' where a': "a' ∈ closed_segment b c" and "d ∙ a' = k"
     and ba': "⋀z. z ∈ closed_segment b a' ⟹ d ∙ z ≤ k"
     and a'c: "⋀z. z ∈ closed_segment a' c ⟹ k ≤ d ∙ z"
    apply (rule hol_pal_lem0 [of d b k c, OF ‹d ∙ b ≤ k›])
    using False by auto
  obtain b' where b': "b' ∈ closed_segment a c" and "d ∙ b' = k"
     and ab': "⋀z. z ∈ closed_segment a b' ⟹ d ∙ z ≤ k"
     and b'c: "⋀z. z ∈ closed_segment b' c ⟹ k ≤ d ∙ z"
    apply (rule hol_pal_lem0 [of d a k c, OF ‹d ∙ a ≤ k›])
    using False by auto
  have a'b': "a' ∈ S ∧ b' ∈ S"
    using a' abc b' convex_contains_segment ‹convex S› by auto
  have "continuous_on (closed_segment c a) f"
    by (meson abc contf continuous_on_subset convex_contains_segment ‹convex S›)
  then have 1: "contour_integral (linepath c a) f =
                contour_integral (linepath c b') f + contour_integral (linepath b' a) f"
    apply (rule contour_integral_split_linepath)
    using b' by (simp add: closed_segment_commute)
  have "continuous_on (closed_segment b c) f"
    by (meson abc contf continuous_on_subset convex_contains_segment ‹convex S›)
  then have 2: "contour_integral (linepath b c) f =
                contour_integral (linepath b a') f + contour_integral (linepath a' c) f"
    by (rule contour_integral_split_linepath [OF _ a'])
  have 3: "contour_integral (reversepath (linepath b' a')) f =
                - contour_integral (linepath b' a') f"
    by (rule contour_integral_reversepath [OF valid_path_linepath])
  have fcd_le: "f field_differentiable at x"
               if "x ∈ interior S ∧ x ∈ interior {x. d ∙ x ≤ k}" for x
  proof -
    have "f holomorphic_on S ∩ {c. d ∙ c < k}"
      by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1)
    then have "∃C D. x ∈ interior C ∩ interior D ∧ f holomorphic_on interior C ∩ interior D"
      using that
      by (metis Collect_mem_eq Int_Collect ‹d ≠ 0› interior_halfspace_le interior_open ‹open S›)
    then show "f field_differentiable at x"
      by (metis at_within_interior holomorphic_on_def interior_Int interior_interior)
  qed
  have ab_le: "⋀x. x ∈ closed_segment a b ⟹ d ∙ x ≤ k"
  proof -
    fix x :: complex
    assume "x ∈ closed_segment a b"
    then have "⋀C. x ∈ C ∨ b ∉ C ∨ a ∉ C ∨ ¬ convex C"
      by (meson contra_subsetD convex_contains_segment)
    then show "d ∙ x ≤ k"
      by (metis lek convex_halfspace_le mem_Collect_eq)
  qed
  have "continuous_on (S ∩ {x. d ∙ x ≤ k}) f" using contf
    by (simp add: continuous_on_subset)
  then have "(f has_contour_integral 0)
         (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)"
    apply (rule Cauchy_theorem_convex [where k = "{}"])
    apply (simp_all add: path_image_join convex_Int convex_halfspace_le ‹convex S› fcd_le ab_le
                closed_segment_subset abc a'b' ba')
    by (metis ‹d ∙ a' = k› ‹d ∙ b' = k› convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl)
  then have 4: "contour_integral (linepath a b) f +
                contour_integral (linepath b a') f +
                contour_integral (linepath a' b') f +
                contour_integral (linepath b' a) f = 0"
    by (rule has_chain_integral_chain_integral4)
  have fcd_ge: "f field_differentiable at x"
               if "x ∈ interior S ∧ x ∈ interior {x. k ≤ d ∙ x}" for x
  proof -
    have f2: "f holomorphic_on S ∩ {c. k < d ∙ c}"
      by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2)
    have f3: "interior S = S"
      by (simp add: interior_open ‹open S›)
    then have "x ∈ S ∩ interior {c. k ≤ d ∙ c}"
      using that by simp
    then show "f field_differentiable at x"
      using f3 f2 unfolding holomorphic_on_def
      by (metis (no_types) ‹d ≠ 0› at_within_interior interior_Int interior_halfspace_ge interior_interior)
  qed
  have "continuous_on (S ∩ {x. k ≤ d ∙ x}) f" using contf
    by (simp add: continuous_on_subset)
  then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')"
    apply (rule Cauchy_theorem_convex [where k = "{}"])
    apply (simp_all add: path_image_join convex_Int convex_halfspace_ge ‹convex S›
                      fcd_ge closed_segment_subset abc a'b' a'c)
    by (metis ‹d ∙ a' = k› b'c closed_segment_commute convex_contains_segment
              convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl)
  then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0"
    by (rule has_chain_integral_chain_integral3)
  show ?thesis
    using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath)
qed

lemma hol_pal_lem3:
  assumes S: "convex S" "open S"
      and abc: "a ∈ S" "b ∈ S" "c ∈ S"
      and "d ≠ 0" and lek: "d ∙ a ≤ k"
      and holf1: "f holomorphic_on {z. z ∈ S ∧ d ∙ z < k}"
      and holf2: "f holomorphic_on {z. z ∈ S ∧ k < d ∙ z}"
      and contf: "continuous_on S f"
    shows "contour_integral (linepath a b) f +
           contour_integral (linepath b c) f +
           contour_integral (linepath c a) f = 0"
proof (cases "d ∙ b ≤ k")
  case True show ?thesis
    by (rule hol_pal_lem2 [OF S abc ‹d ≠ 0› lek True holf1 holf2 contf])
next
  case False
  show ?thesis
  proof (cases "d ∙ c ≤ k")
    case True
    have "contour_integral (linepath c a) f +
          contour_integral (linepath a b) f +
          contour_integral (linepath b c) f = 0"
      by (rule hol_pal_lem2 [OF S ‹c ∈ S› ‹a ∈ S› ‹b ∈ S› ‹d ≠ 0› ‹d ∙ c ≤ k› lek holf1 holf2 contf])
    then show ?thesis
      by (simp add: algebra_simps)
  next
    case False
    have "contour_integral (linepath b c) f +
          contour_integral (linepath c a) f +
          contour_integral (linepath a b) f = 0"
      apply (rule hol_pal_lem2 [OF S ‹b ∈ S› ‹c ∈ S› ‹a ∈ S›, of "-d" "-k"])
      using ‹d ≠ 0› ‹¬ d ∙ b ≤ k› False by (simp_all add: holf1 holf2 contf)
    then show ?thesis
      by (simp add: algebra_simps)
  qed
qed

lemma hol_pal_lem4:
  assumes S: "convex S" "open S"
      and abc: "a ∈ S" "b ∈ S" "c ∈ S" and "d ≠ 0"
      and holf1: "f holomorphic_on {z. z ∈ S ∧ d ∙ z < k}"
      and holf2: "f holomorphic_on {z. z ∈ S ∧ k < d ∙ z}"
      and contf: "continuous_on S f"
    shows "contour_integral (linepath a b) f +
           contour_integral (linepath b c) f +
           contour_integral (linepath c a) f = 0"
proof (cases "d ∙ a ≤ k")
  case True show ?thesis
    by (rule hol_pal_lem3 [OF S abc ‹d ≠ 0› True holf1 holf2 contf])
next
  case False
  show ?thesis
    apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"])
    using ‹d ≠ 0› False by (simp_all add: holf1 holf2 contf)
qed

proposition holomorphic_on_paste_across_line:
  assumes S: "open S" and "d ≠ 0"
      and holf1: "f holomorphic_on (S ∩ {z. d ∙ z < k})"
      and holf2: "f holomorphic_on (S ∩ {z. k < d ∙ z})"
      and contf: "continuous_on S f"
    shows "f holomorphic_on S"
proof -
  have *: "∃t. open t ∧ p ∈ t ∧ continuous_on t f ∧
               (∀a b c. convex hull {a, b, c} ⊆ t ⟶
                         contour_integral (linepath a b) f +
                         contour_integral (linepath b c) f +
                         contour_integral (linepath c a) f = 0)"
          if "p ∈ S" for p
  proof -
    obtain e where "e>0" and e: "ball p e ⊆ S"
      using ‹p ∈ S› openE S by blast
    then have "continuous_on (ball p e) f"
      using contf continuous_on_subset by blast
    moreover have "f holomorphic_on {z. dist p z < e ∧ d ∙ z < k}"
      apply (rule holomorphic_on_subset [OF holf1])
      using e by auto
    moreover have "f holomorphic_on {z. dist p z < e ∧ k < d ∙ z}"
      apply (rule holomorphic_on_subset [OF holf2])
      using e by auto
    ultimately show ?thesis
      apply (rule_tac x="ball p e" in exI)
      using ‹e > 0› e ‹d ≠ 0›
      apply (simp add:, clarify)
      apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k])
      apply (auto simp: subset_hull)
      done
  qed
  show ?thesis
    by (blast intro: * Morera_local_triangle analytic_imp_holomorphic)
qed

proposition Schwarz_reflection:
  assumes "open S" and cnjs: "cnj ` S ⊆ S"
      and  holf: "f holomorphic_on (S ∩ {z. 0 < Im z})"
      and contf: "continuous_on (S ∩ {z. 0 ≤ Im z}) f"
      and f: "⋀z. ⟦z ∈ S; z ∈ ℝ⟧ ⟹ (f z) ∈ ℝ"
    shows "(λz. if 0 ≤ Im z then f z else cnj(f(cnj z))) holomorphic_on S"
proof -
  have 1: "(λz. if 0 ≤ Im z then f z else cnj (f (cnj z))) holomorphic_on (S ∩ {z. 0 < Im z})"
    by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf])
  have cont_cfc: "continuous_on (S ∩ {z. Im z ≤ 0}) (cnj o f o cnj)"
    apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf])
    using cnjs apply auto
    done
  have "cnj ∘ f ∘ cnj field_differentiable at x within S ∩ {z. Im z < 0}"
        if "x ∈ S" "Im x < 0" "f field_differentiable at (cnj x) within S ∩ {z. 0 < Im z}" for x
    using that
    apply (simp add: field_differentiable_def Derivative.DERIV_within_iff Lim_within dist_norm, clarify)
    apply (rule_tac x="cnj f'" in exI)
    apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify)
    apply (drule_tac x="cnj xa" in bspec)
    using cnjs apply force
    apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj)
    done
  then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S ∩ {z. Im z < 0})"
    using holf cnjs
    by (force simp: holomorphic_on_def)
  have 2: "(λz. if 0 ≤ Im z then f z else cnj (f (cnj z))) holomorphic_on (S ∩ {z. Im z < 0})"
    apply (rule iffD1 [OF holomorphic_cong [OF refl]])
    using hol_cfc by auto
  have [simp]: "(S ∩ {z. 0 ≤ Im z}) ∪ (S ∩ {z. Im z ≤ 0}) = S"
    by force
  have "continuous_on ((S ∩ {z. 0 ≤ Im z}) ∪ (S ∩ {z. Im z ≤ 0}))
                       (λz. if 0 ≤ Im z then f z else cnj (f (cnj z)))"
    apply (rule continuous_on_cases_local)
    using cont_cfc contf
    apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge)
    using f Reals_cnj_iff complex_is_Real_iff apply auto
    done
  then have 3: "continuous_on S (λz. if 0 ≤ Im z then f z else cnj (f (cnj z)))"
    by force
  show ?thesis
    apply (rule holomorphic_on_paste_across_line [OF ‹open S›, of "- 𝗂" _ 0])
    using 1 2 3
    apply auto
    done
qed

subsection‹Bloch's theorem›

lemma Bloch_lemma_0:
  assumes holf: "f holomorphic_on cball 0 r" and "0 < r"
      and [simp]: "f 0 = 0"
      and le: "⋀z. norm z < r ⟹ norm(deriv f z) ≤ 2 * norm(deriv f 0)"
    shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) ⊆ f ` ball 0 r"
proof -
  have "sqrt 2 < 3/2"
    by (rule real_less_lsqrt) (auto simp: power2_eq_square)
  then have sq3: "0 < 3 - 2 * sqrt 2" by simp
  show ?thesis
  proof (cases "deriv f 0 = 0")
    case True then show ?thesis by simp
  next
    case False
    define C where "C = 2 * norm(deriv f 0)"
    have "0 < C" using False by (simp add: C_def)
    have holf': "f holomorphic_on ball 0 r" using holf
      using ball_subset_cball holomorphic_on_subset by blast
    then have holdf': "deriv f holomorphic_on ball 0 r"
      by (rule holomorphic_deriv [OF _ open_ball])
    have "Le1": "norm(deriv f z - deriv f 0) ≤ norm z / (r - norm z) * C"
                if "norm z < r" for z
    proof -
      have T1: "norm(deriv f z - deriv f 0) ≤ norm z / (R - norm z) * C"
              if R: "norm z < R" "R < r" for R
      proof -
        have "0 < R" using R
          by (metis less_trans norm_zero zero_less_norm_iff)
        have df_le: "⋀x. norm x < r ⟹ norm (deriv f x) ≤ C"
          using le by (simp add: C_def)
        have hol_df: "deriv f holomorphic_on cball 0 R"
          apply (rule holomorphic_on_subset) using R holdf' by auto
        have *: "((λw. deriv f w / (w - z)) has_contour_integral 2 * pi * 𝗂 * deriv f z) (circlepath 0 R)"
                 if "norm z < R" for z
          using ‹0 < R› that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"]
          by (force simp: winding_number_circlepath)
        have **: "((λx. deriv f x / (x - z) - deriv f x / x) has_contour_integral
                   of_real (2 * pi) * 𝗂 * (deriv f z - deriv f 0))
                  (circlepath 0 R)"
           using has_contour_integral_diff [OF * [of z] * [of 0]] ‹0 < R› that
           by (simp add: algebra_simps)
        have [simp]: "⋀x. norm x = R ⟹ x ≠ z"  using that(1) by blast
        have "norm (deriv f x / (x - z) - deriv f x / x)
                     ≤ C * norm z / (R * (R - norm z))"
                  if "norm x = R" for x
        proof -
          have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) =
                        norm (deriv f x) * norm z"
            by (simp add: norm_mult right_diff_distrib')
          show ?thesis
            using  ‹0 < R› ‹0 < C› R that
            apply (simp add: norm_mult norm_divide divide_simps)
            using df_le norm_triangle_ineq2 ‹0 < C› apply (auto intro!: mult_mono)
            done
        qed
        then show ?thesis
          using has_contour_integral_bound_circlepath
                  [OF **, of "C * norm z/(R*(R - norm z))"]
                ‹0 < R› ‹0 < C› R
          apply (simp add: norm_mult norm_divide)
          apply (simp add: divide_simps mult.commute)
          done
      qed
      obtain r' where r': "norm z < r'" "r' < r"
        using Rats_dense_in_real [of "norm z" r] ‹norm z < r› by blast
      then have [simp]: "closure {r'<..<r} = {r'..r}" by simp
      show ?thesis
        apply (rule continuous_ge_on_closure
                 [where f = "λr. norm z / (r - norm z) * C" and s = "{r'<..<r}",
                  OF _ _ T1])
        apply (intro continuous_intros)
        using that r'
        apply (auto simp: not_le)
        done
    qed
    have "*": "(norm z - norm z^2/(r - norm z)) * norm(deriv f 0) ≤ norm(f z)"
              if r: "norm z < r" for z
    proof -
      have 1: "⋀x. x ∈ ball 0 r ⟹
              ((λz. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0)
               (at x within ball 0 r)"
        by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+
      have 2: "closed_segment 0 z ⊆ ball 0 r"
        by (metis ‹0 < r› convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that)
      have 3: "(λt. (norm z)2 * t / (r - norm z) * C) integrable_on {0..1}"
        apply (rule integrable_on_cmult_right [where 'b=real, simplified])
        apply (rule integrable_on_cdivide [where 'b=real, simplified])
        apply (rule integrable_on_cmult_left [where 'b=real, simplified])
        apply (rule ident_integrable_on)
        done
      have 4: "norm (deriv f (x *R z) - deriv f 0) * norm z ≤ norm z * norm z * x * C / (r - norm z)"
              if x: "0 ≤ x" "x ≤ 1" for x
      proof -
        have [simp]: "x * norm z < r"
          using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero)
        have "norm (deriv f (x *R z) - deriv f 0) ≤ norm (x *R z) / (r - norm (x *R z)) * C"
          apply (rule Le1) using r x ‹0 < r› by simp
        also have "... ≤ norm (x *R z) / (r - norm z) * C"
          using r x ‹0 < r›
          apply (simp add: divide_simps)
          by (simp add: ‹0 < C› mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono)
        finally have "norm (deriv f (x *R z) - deriv f 0) * norm z ≤ norm (x *R z)  / (r - norm z) * C * norm z"
          by (rule mult_right_mono) simp
        with x show ?thesis by (simp add: algebra_simps)
      qed
      have le_norm: "abc ≤ norm d - e ⟹ norm(f - d) ≤ e ⟹ abc ≤ norm f" for abc d e and f::complex
        by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans)
      have "norm (integral {0..1} (λx. (deriv f (x *R z) - deriv f 0) * z))
            ≤ integral {0..1} (λt. (norm z)2 * t / (r - norm z) * C)"
        apply (rule integral_norm_bound_integral)
        using contour_integral_primitive [OF 1, of "linepath 0 z"] 2
        apply (simp add: has_contour_integral_linepath has_integral_integrable_integral)
        apply (rule 3)
        apply (simp add: norm_mult power2_eq_square 4)
        done
      then have int_le: "norm (f z - deriv f 0 * z) ≤ (norm z)2 * norm(deriv f 0) / ((r - norm z))"
        using contour_integral_primitive [OF 1, of "linepath 0 z"] 2
        apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def)
        done
      show ?thesis
        apply (rule le_norm [OF _ int_le])
        using ‹norm z < r›
        apply (simp add: power2_eq_square divide_simps C_def norm_mult)
        proof -
          have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) ≤ norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)"
            by (simp add: linordered_field_class.sign_simps(38))
          then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) ≤ norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)"
            by (simp add: linordered_field_class.sign_simps(38) mult.commute mult.left_commute)
        qed
    qed
    have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2)  < 1"
      by (auto simp:  sqrt2_less_2)
    have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f"
      apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]])
      apply (subst closure_ball)
      using ‹0 < r› mult_pos_pos sq201
      apply (auto simp: cball_subset_cball_iff)
      done
    have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))"
      apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force)
      using ‹0 < r› mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff)
      using False ‹0 < r› centre_in_ball holf' holomorphic_nonconstant by blast
    have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) =
          ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))"
      by simp
    also have "...  ⊆ f ` ball 0 ((1 - sqrt 2 / 2) * r)"
    proof -
      have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) ≤ norm (f z)"
           if "norm z = (1 - sqrt 2 / 2) * r" for z
        apply (rule order_trans [OF _ *])
        using  ‹0 < r›
        apply (simp_all add: field_simps  power2_eq_square that)
        apply (simp add: mult.assoc [symmetric])
        done
      show ?thesis
        apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball])
        using ‹0 < r› sq201 3 apply simp_all
        using C_def ‹0 < C› sq3 apply force
        done
     qed
    also have "...  ⊆ f ` ball 0 r"
      apply (rule image_subsetI [OF imageI], simp)
      apply (erule less_le_trans)
      using ‹0 < r› apply (auto simp: field_simps)
      done
    finally show ?thesis .
  qed
qed

lemma Bloch_lemma:
  assumes holf: "f holomorphic_on cball a r" and "0 < r"
      and le: "⋀z. z ∈ ball a r ⟹ norm(deriv f z) ≤ 2 * norm(deriv f a)"
    shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) ⊆ f ` ball a r"
proof -
  have fz: "(λz. f (a + z)) = f o (λz. (a + z))"
    by (simp add: o_def)
  have hol0: "(λz. f (a + z)) holomorphic_on cball 0 r"
    unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+
  then have [simp]: "⋀x. norm x < r ⟹ (λz. f (a + z)) field_differentiable at x"
    by (metis Topology_Euclidean_Space.open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel)
  have [simp]: "⋀z. norm z < r ⟹ f field_differentiable at (a + z)"
    by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute)
  then have [simp]: "f field_differentiable at a"
    by (metis add.comm_neutral ‹0 < r› norm_eq_zero)
  have hol1: "(λz. f (a + z) - f a) holomorphic_on cball 0 r"
    by (intro holomorphic_intros hol0)
  then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (λz. f (a + z) - f a) 0))
             ⊆ (λz. f (a + z) - f a) ` ball 0 r"
    apply (rule Bloch_lemma_0)
    apply (simp_all add: ‹0 < r›)
    apply (simp add: fz complex_derivative_chain)
    apply (simp add: dist_norm le)
    done
  then show ?thesis
    apply clarify
    apply (drule_tac c="x - f a" in subsetD)
     apply (force simp: fz ‹0 < r› dist_norm complex_derivative_chain field_differentiable_compose)+
    done
qed

proposition Bloch_unit:
  assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1"
  obtains b r where "1/12 < r" "ball b r ⊆ f ` (ball a 1)"
proof -
  define r :: real where "r = 249/256"
  have "0 < r" "r < 1" by (auto simp: r_def)
  define g where "g z = deriv f z * of_real(r - norm(z - a))" for z
  have "deriv f holomorphic_on ball a 1"
    by (rule holomorphic_deriv [OF holf open_ball])
  then have "continuous_on (ball a 1) (deriv f)"
    using holomorphic_on_imp_continuous_on by blast
  then have "continuous_on (cball a r) (deriv f)"
    by (rule continuous_on_subset) (simp add: cball_subset_ball_iff ‹r < 1›)
  then have "continuous_on (cball a r) g"
    by (simp add: g_def continuous_intros)
  then have 1: "compact (g ` cball a r)"
    by (rule compact_continuous_image [OF _ compact_cball])
  have 2: "g ` cball a r ≠ {}"
    using ‹r > 0› by auto
  obtain p where pr: "p ∈ cball a r"
             and pge: "⋀y. y ∈ cball a r ⟹ norm (g y) ≤ norm (g p)"
    using distance_attains_sup [OF 1 2, of 0] by force
  define t where "t = (r - norm(p - a)) / 2"
  have "norm (p - a) ≠ r"
    using pge [of a] ‹r > 0› by (auto simp: g_def norm_mult)
  then have "norm (p - a) < r" using pr
    by (simp add: norm_minus_commute dist_norm)
  then have "0 < t"
    by (simp add: t_def)
  have cpt: "cball p t ⊆ ball a r"
    using ‹0 < t› by (simp add: cball_subset_ball_iff dist_norm t_def field_simps)
  have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) ≤ norm (deriv f p)"
            if "y ∈ cball a r" for y
  proof -
    have [simp]: "norm (y - a) ≤ r"
      using that by (simp add: dist_norm norm_minus_commute)
    have "norm (g y) ≤ norm (g p)"
      using pge [OF that] by simp
    then have "norm (deriv f y) * abs (r - norm (y - a)) ≤ norm (deriv f p) * abs (r - norm (p - a))"
      by (simp only: dist_norm g_def norm_mult norm_of_real)
    with that ‹norm (p - a) < r› show ?thesis
      by (simp add: dist_norm divide_simps)
  qed
  have le_norm_dfp: "r / (r - norm (p - a)) ≤ norm (deriv f p)"
    using gen_le_dfp [of a] ‹r > 0› by auto
  have 1: "f holomorphic_on cball p t"
    apply (rule holomorphic_on_subset [OF holf])
    using cpt ‹r < 1› order_subst1 subset_ball by auto
  have 2: "norm (deriv f z) ≤ 2 * norm (deriv f p)" if "z ∈ ball p t" for z
  proof -
    have z: "z ∈ cball a r"
      by (meson ball_subset_cball subsetD cpt that)
    then have "norm(z - a) < r"
      by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that)
    have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) ≤ norm (deriv f p)"
      using gen_le_dfp [OF z] by simp
    with ‹norm (z - a) < r› ‹norm (p - a) < r›
    have "norm (deriv f z) ≤ (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)"
       by (simp add: field_simps)
    also have "... ≤ 2 * norm (deriv f p)"
      apply (rule mult_right_mono)
      using that ‹norm (p - a) < r› ‹norm(z - a) < r›
      apply (simp_all add: field_simps t_def dist_norm [symmetric])
      using dist_triangle3 [of z a p] by linarith
    finally show ?thesis .
  qed
  have sqrt2: "sqrt 2 < 2113/1494"
    by (rule real_less_lsqrt) (auto simp: power2_eq_square)
  then have sq3: "0 < 3 - 2 * sqrt 2" by simp
  have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r"
    using sq3 sqrt2 by (auto simp: field_simps r_def)
  also have "... ≤ cmod (deriv f p) * (r - cmod (p - a))"
    using ‹norm (p - a) < r› le_norm_dfp   by (simp add: pos_divide_le_eq)
  finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)"
    using pos_divide_less_eq half_gt_zero_iff sq3 by blast
  then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)"
    using sq3 by (simp add: mult.commute t_def)
  have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) ⊆ f ` ball p t"
    by (rule Bloch_lemma [OF 1 ‹0 < t› 2])
  also have "... ⊆ f ` ball a 1"
    apply (rule image_mono)
    apply (rule order_trans [OF ball_subset_cball])
    apply (rule order_trans [OF cpt])
    using ‹0 < t› ‹r < 1› apply (simp add: ball_subset_ball_iff dist_norm)
    done
  finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) ⊆ f ` ball a 1" .
  with ** show ?thesis
    by (rule that)
qed


theorem Bloch:
  assumes holf: "f holomorphic_on ball a r" and "0 < r"
      and r': "r' ≤ r * norm (deriv f a) / 12"
  obtains b where "ball b r' ⊆ f ` (ball a r)"
proof (cases "deriv f a = 0")
  case True with r' show ?thesis
    using ball_eq_empty that by fastforce
next
  case False
  define C where "C = deriv f a"
  have "0 < norm C" using False by (simp add: C_def)
  have dfa: "f field_differentiable at a"
    apply (rule holomorphic_on_imp_differentiable_at [OF holf])
    using ‹0 < r› by auto
  have fo: "(λz. f (a + of_real r * z)) = f o (λz. (a + of_real r * z))"
    by (simp add: o_def)
  have holf': "f holomorphic_on (λz. a + complex_of_real r * z) ` ball 0 1"
    apply (rule holomorphic_on_subset [OF holf])
    using ‹0 < r› apply (force simp: dist_norm norm_mult)
    done
  have 1: "(λz. f (a + r * z) / (C * r)) holomorphic_on ball 0 1"
    apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+
    using ‹0 < r› by (simp add: C_def False)
  have "((λz. f (a + of_real r * z) / (C * of_real r)) has_field_derivative
        (deriv f (a + of_real r * z) / C)) (at z)"
       if "norm z < 1" for z
  proof -
    have *: "((λx. f (a + of_real r * x)) has_field_derivative
           (deriv f (a + of_real r * z) * of_real r)) (at z)"
      apply (simp add: fo)
      apply (rule DERIV_chain [OF field_differentiable_derivI])
      apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp)
      using ‹0 < r› apply (simp add: dist_norm norm_mult that)
      apply (rule derivative_eq_intros | simp)+
      done
    show ?thesis
      apply (rule derivative_eq_intros * | simp)+
      using ‹0 < r› by (auto simp: C_def False)
  qed
  have 2: "deriv (λz. f (a + of_real r * z) / (C * of_real r)) 0 = 1"
    apply (subst deriv_cdivide_right)
    apply (simp add: field_differentiable_def fo)
    apply (rule exI)
    apply (rule DERIV_chain [OF field_differentiable_derivI])
    apply (simp add: dfa)
    apply (rule derivative_eq_intros | simp add: C_def False fo)+
    using ‹0 < r›
    apply (simp add: C_def False fo)
    apply (simp add: derivative_intros dfa complex_derivative_chain)
    done
  have sb1: "op * (C * r) ` (λz. f (a + of_real r * z) / (C * r)) ` ball 0 1
             ⊆ f ` ball a r"
    using ‹0 < r› by (auto simp: dist_norm norm_mult C_def False)
  have sb2: "ball (C * r * b) r' ⊆ op * (C * r) ` ball b t"
             if "1 / 12 < t" for b t
  proof -
    have *: "r * cmod (deriv f a) / 12 ≤ r * (t * cmod (deriv f a))"
      using that ‹0 < r› less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right
      by auto
    show ?thesis
      apply clarify
      apply (rule_tac x="x / (C * r)" in image_eqI)
      using ‹0 < r›
      apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps)
      apply (erule less_le_trans)
      apply (rule order_trans [OF r' *])
      done
  qed
  show ?thesis
    apply (rule Bloch_unit [OF 1 2])
    apply (rename_tac t)
    apply (rule_tac b="(C * of_real r) * b" in that)
    apply (drule image_mono [where f = "λz. (C * of_real r) * z"])
    using sb1 sb2
    apply force
    done
qed

corollary Bloch_general:
  assumes holf: "f holomorphic_on s" and "a ∈ s"
      and tle: "⋀z. z ∈ frontier s ⟹ t ≤ dist a z"
      and rle: "r ≤ t * norm(deriv f a) / 12"
  obtains b where "ball b r ⊆ f ` s"
proof -
  consider "r ≤ 0" | "0 < t * norm(deriv f a) / 12" using rle by force
  then show ?thesis
  proof cases
    case 1 then show ?thesis
      by (simp add: Topology_Euclidean_Space.ball_empty that)
  next
    case 2
    show ?thesis
    proof (cases "deriv f a = 0")
      case True then show ?thesis
        using rle by (simp add: Topology_Euclidean_Space.ball_empty that)
    next
      case False
      then have "t > 0"
        using 2 by (force simp: zero_less_mult_iff)
      have "~ ball a t ⊆ s ⟹ ball a t ∩ frontier s ≠ {}"
        apply (rule connected_Int_frontier [of "ball a t" s], simp_all)
        using ‹0 < t› ‹a ∈ s› centre_in_ball apply blast
        done
      with tle have *: "ball a t ⊆ s" by fastforce
      then have 1: "f holomorphic_on ball a t"
        using holf using holomorphic_on_subset by blast
      show ?thesis
        apply (rule Bloch [OF 1 ‹t > 0› rle])
        apply (rule_tac b=b in that)
        using * apply force
        done
    qed
  qed
qed

subsection ‹Foundations of Cauchy's residue theorem›

text‹Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem.
    Interactive Theorem Proving›

definition residue :: "(complex ⇒ complex) ⇒ complex ⇒ complex" where
  "residue f z = (SOME int. ∃e>0. ∀ε>0. ε<e
    ⟶ (f has_contour_integral 2*pi* 𝗂 *int) (circlepath z ε))"

lemma Eps_cong:
  assumes "⋀x. P x = Q x"
  shows   "Eps P = Eps Q"
  using ext[of P Q, OF assms] by simp

lemma residue_cong:
  assumes eq: "eventually (λz. f z = g z) (at z)" and "z = z'"
  shows   "residue f z = residue g z'"
proof -
  from assms have eq': "eventually (λz. g z = f z) (at z)"
    by (simp add: eq_commute)
  let ?P = "λf c e. (∀ε>0. ε < e ⟶
   (f has_contour_integral of_real (2 * pi) * 𝗂 * c) (circlepath z ε))"
  have "residue f z = residue g z" unfolding residue_def
  proof (rule Eps_cong)
    fix c :: complex
    have "∃e>0. ?P g c e" 
      if "∃e>0. ?P f c e" and "eventually (λz. f z = g z) (at z)" for f g 
    proof -
      from that(1) obtain e where e: "e > 0" "?P f c e"
        by blast
      from that(2) obtain e' where e': "e' > 0" "⋀z'. z' ≠ z ⟹ dist z' z < e' ⟹ f z' = g z'"
        unfolding eventually_at by blast
      have "?P g c (min e e')"
      proof (intro allI exI impI, goal_cases)
        case (1 ε)
        hence "(f has_contour_integral of_real (2 * pi) * 𝗂 * c) (circlepath z ε)" 
          using e(2) by auto
        thus ?case
        proof (rule has_contour_integral_eq)
          fix z' assume "z' ∈ path_image (circlepath z ε)"
          hence "dist z' z < e'" and "z' ≠ z"
            using 1 by (auto simp: dist_commute)
          with e'(2)[of z'] show "f z' = g z'" by simp
        qed
      qed
      moreover from e and e' have "min e e' > 0" by auto
      ultimately show ?thesis by blast
    qed
    from this[OF _ eq] and this[OF _ eq']
      show "(∃e>0. ?P f c e) ⟷ (∃e>0. ?P g c e)"
      by blast
  qed
  with assms show ?thesis by simp
qed

lemma contour_integral_circlepath_eq:
  assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0<e1" "e1≤e2"
    and e2_cball:"cball z e2 ⊆ s"
  shows
    "f contour_integrable_on circlepath z e1"
    "f contour_integrable_on circlepath z e2"
    "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f"
proof -
  define l where "l ≡ linepath (z+e2) (z+e1)"
  have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto
  have "e2>0" using ‹e1>0› ‹e1≤e2› by auto
  have zl_img:"z∉path_image l"
    proof
      assume "z ∈ path_image l"
      then have "e2 ≤ cmod (e2 - e1)"
        using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] ‹e1>0› ‹e2>0› unfolding l_def
        by (auto simp add:closed_segment_commute)
      thus False using ‹e2>0› ‹e1>0› ‹e1≤e2›
        apply (subst (asm) norm_of_real)
        by auto
    qed
  define g where "g ≡ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l"
  show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)"
    proof -
      show "f contour_integrable_on circlepath z e2"
        apply (intro contour_integrable_continuous_circlepath[OF
                continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
        using ‹e2>0› e2_cball by auto
      show "f contour_integrable_on (circlepath z e1)"
        apply (intro contour_integrable_continuous_circlepath[OF
                      continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
        using ‹e1>0› ‹e1≤e2› e2_cball by auto
    qed
  have [simp]:"f contour_integrable_on l"
    proof -
      have "closed_segment (z + e2) (z + e1) ⊆ cball z e2" using ‹e2>0› ‹e1>0› ‹e1≤e2›
        by (intro closed_segment_subset,auto simp add:dist_norm)
      hence "closed_segment (z + e2) (z + e1) ⊆ s - {z}" using zl_img e2_cball unfolding l_def
        by auto
      then show "f contour_integrable_on l" unfolding l_def
        apply (intro contour_integrable_continuous_linepath[OF
                      continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
        by auto
    qed
  let ?ig="λg. contour_integral g f"
  have "(f has_contour_integral 0) g"
    proof (rule Cauchy_theorem_global[OF _ f_holo])
      show "open (s - {z})" using ‹open s› by auto
      show "valid_path g" unfolding g_def l_def by auto
      show "pathfinish g = pathstart g" unfolding g_def l_def by auto
    next
      have path_img:"path_image g ⊆ cball z e2"
        proof -
          have "closed_segment (z + e2) (z + e1) ⊆ cball z e2" using ‹e2>0› ‹e1>0› ‹e1≤e2›
            by (intro closed_segment_subset,auto simp add:dist_norm)
          moreover have "sphere z ¦e1¦ ⊆ cball z e2" using ‹e2>0› ‹e1≤e2› ‹e1>0› by auto
          ultimately show ?thesis unfolding g_def l_def using ‹e2>0›
            by (simp add: path_image_join closed_segment_commute)
        qed
      show "path_image g ⊆ s - {z}"
        proof -
          have "z∉path_image g" using zl_img
            unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute)
          moreover note ‹cball z e2 ⊆ s› and path_img
          ultimately show ?thesis by auto
        qed
      show "winding_number g w = 0" when"w ∉ s - {z}" for w
        proof -
          have "winding_number g w = 0" when "w∉s" using that e2_cball
            apply (intro winding_number_zero_outside[OF _ _ _ _ path_img])
            by (auto simp add:g_def l_def)
          moreover have "winding_number g z=0"
            proof -
              let ?Wz="λg. winding_number g z"
              have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1))
                  + ?Wz (reversepath l)"
                using ‹e2>0› ‹e1>0› zl_img unfolding g_def l_def
                by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+
              also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))"
                using zl_img
                apply (subst (2) winding_number_reversepath)
                by (auto simp add:l_def closed_segment_commute)
              also have "... = 0"
                proof -
                  have "?Wz (circlepath z e2) = 1" using ‹e2>0›
                    by (auto intro: winding_number_circlepath_centre)
                  moreover have "?Wz (reversepath (circlepath z e1)) = -1" using ‹e1>0›
                    apply (subst winding_number_reversepath)
                    by (auto intro: winding_number_circlepath_centre)
                  ultimately show ?thesis by auto
                qed
              finally show ?thesis .
            qed
          ultimately show ?thesis using that by auto
        qed
    qed
  then have "0 = ?ig g" using contour_integral_unique by simp
  also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1))
      + ?ig (reversepath l)"
    unfolding g_def
    by (auto simp add:contour_integrable_reversepath_eq)
  also have "... = ?ig (circlepath z e2)  - ?ig (circlepath z e1)"
    by (auto simp add:contour_integral_reversepath)
  finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f"
    by simp
qed

lemma base_residue:
  assumes "open s" "z∈s" "r>0" and f_holo:"f holomorphic_on (s - {z})"
    and r_cball:"cball z r ⊆ s"
  shows "(f has_contour_integral 2 * pi * 𝗂 * (residue f z)) (circlepath z r)"
proof -
  obtain e where "e>0" and e_cball:"cball z e ⊆ s"
    using open_contains_cball[of s] ‹open s› ‹z∈s› by auto
  define c where "c ≡ 2 * pi * 𝗂"
  define i where "i ≡ contour_integral (circlepath z e) f / c"
  have "(f has_contour_integral c*i) (circlepath z ε)" when "ε>0" "ε<e" for ε
    proof -
      have "contour_integral (circlepath z e) f = contour_integral (circlepath z ε) f"
          "f contour_integrable_on circlepath z ε"
          "f contour_integrable_on circlepath z e"
        using ‹ε<e›
        by (intro contour_integral_circlepath_eq[OF ‹open s› f_holo ‹ε>0› _ e_cball],auto)+
      then show ?thesis unfolding i_def c_def
        by (auto intro:has_contour_integral_integral)
    qed
  then have "∃e>0. ∀ε>0. ε<e ⟶ (f has_contour_integral c * (residue f z)) (circlepath z ε)"
    unfolding residue_def c_def
    apply (rule_tac someI[of _ i],intro  exI[where x=e])
    by (auto simp add:‹e>0› c_def)
  then obtain e' where "e'>0"
      and e'_def:"∀ε>0. ε<e' ⟶ (f has_contour_integral c * (residue f z)) (circlepath z ε)"
    by auto
  let ?int="λe. contour_integral (circlepath z e) f"
  define  ε where "ε ≡ Min {r,e'} / 2"
  have "ε>0" "ε≤r" "ε<e'" using ‹r>0› ‹e'>0› unfolding ε_def by auto
  have "(f has_contour_integral c * (residue f z)) (circlepath z ε)"
    using e'_def[rule_format,OF ‹ε>0› ‹ε<e'›] .
  then show ?thesis unfolding c_def
    using contour_integral_circlepath_eq[OF ‹open s› f_holo ‹ε>0› ‹ε≤r› r_cball]
    by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z ε" "circlepath z r"])
qed


lemma residue_holo:
  assumes "open s" "z ∈ s" and f_holo: "f holomorphic_on s"
  shows "residue f z = 0"
proof -
  define c where "c ≡ 2 * pi * 𝗂"
  obtain e where "e>0" and e_cball:"cball z e ⊆ s" using ‹open s› ‹z∈s›
    using open_contains_cball_eq by blast
  have "(f has_contour_integral c*residue f z) (circlepath z e)"
    using f_holo
    by (auto intro: base_residue[OF ‹open s› ‹z∈s› ‹e>0› _ e_cball,folded c_def])
  moreover have "(f has_contour_integral 0) (circlepath z e)"
    using f_holo e_cball ‹e>0›
    by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"])
  ultimately have "c*residue f z =0"
    using has_contour_integral_unique by blast
  thus ?thesis unfolding c_def  by auto
qed


lemma residue_const:"residue (λ_. c) z = 0"
  by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros)


lemma residue_add:
  assumes "open s" "z ∈ s" and f_holo: "f holomorphic_on s - {z}"
      and g_holo:"g holomorphic_on s - {z}"
  shows "residue (λz. f z + g z) z= residue f z + residue g z"
proof -
  define c where "c ≡ 2 * pi * 𝗂"
  define fg where "fg ≡ (λz. f z+g z)"
  obtain e where "e>0" and e_cball:"cball z e ⊆ s" using ‹open s› ‹z∈s›
    using open_contains_cball_eq by blast
  have "(fg has_contour_integral c * residue fg z) (circlepath z e)"
    unfolding fg_def using f_holo g_holo
    apply (intro base_residue[OF ‹open s› ‹z∈s› ‹e>0› _ e_cball,folded c_def])
    by (auto intro:holomorphic_intros)
  moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)"
    unfolding fg_def using f_holo g_holo
    by (auto intro: has_contour_integral_add base_residue[OF ‹open s› ‹z∈s› ‹e>0› _ e_cball,folded c_def])
  ultimately have "c*(residue f z + residue g z) = c * residue fg z"
    using has_contour_integral_unique by (auto simp add:distrib_left)
  thus ?thesis unfolding fg_def
    by (auto simp add:c_def)
qed


lemma residue_lmul:
  assumes "open s" "z ∈ s" and f_holo: "f holomorphic_on s - {z}"
  shows "residue (λz. c * (f z)) z= c * residue f z"
proof (cases "c=0")
  case True
  thus ?thesis using residue_const by auto
next
  case False
  define c' where "c' ≡ 2 * pi * 𝗂"
  define f' where "f' ≡ (λz. c * (f z))"
  obtain e where "e>0" and e_cball:"cball z e ⊆ s" using ‹open s› ‹z∈s›
    using open_contains_cball_eq by blast
  have "(f' has_contour_integral c' * residue f' z) (circlepath z e)"
    unfolding f'_def using f_holo
    apply (intro base_residue[OF ‹open s› ‹z∈s› ‹e>0› _ e_cball,folded c'_def])
    by (auto intro:holomorphic_intros)
  moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)"
    unfolding f'_def using f_holo
    by (auto intro: has_contour_integral_lmul
      base_residue[OF ‹open s› ‹z∈s› ‹e>0› _ e_cball,folded c'_def])
  ultimately have "c' * residue f' z  = c * (c' * residue f z)"
    using has_contour_integral_unique by auto
  thus ?thesis unfolding f'_def c'_def using False
    by (auto simp add:field_simps)
qed

lemma residue_rmul:
  assumes "open s" "z ∈ s" and f_holo: "f holomorphic_on s - {z}"
  shows "residue (λz. (f z) * c) z= residue f z * c"
using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps)

lemma residue_div:
  assumes "open s" "z ∈ s" and f_holo: "f holomorphic_on s - {z}"
  shows "residue (λz. (f z) / c) z= residue f z / c "
using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps)

lemma residue_neg:
  assumes "open s" "z ∈ s" and f_holo: "f holomorphic_on s - {z}"
  shows "residue (λz. - (f z)) z= - residue f z"
using residue_lmul[OF assms,of "-1"] by auto

lemma residue_diff:
  assumes "open s" "z ∈ s" and f_holo: "f holomorphic_on s - {z}"
      and g_holo:"g holomorphic_on s - {z}"
  shows "residue (λz. f z - g z) z= residue f z - residue g z"
using residue_add[OF assms(1,2,3),of "λz. - g z"] residue_neg[OF assms(1,2,4)]
by (auto intro:holomorphic_intros g_holo)

lemma residue_simple:
  assumes "open s" "z∈s" and f_holo:"f holomorphic_on s"
  shows "residue (λw. f w / (w - z)) z = f z"
proof -
  define c where "c ≡ 2 * pi * 𝗂"
  define f' where "f' ≡ λw. f w / (w - z)"
  obtain e where "e>0" and e_cball:"cball z e ⊆ s" using ‹open s› ‹z∈s›
    using open_contains_cball_eq by blast
  have "(f' has_contour_integral c * f z) (circlepath z e)"
    unfolding f'_def c_def using ‹e>0› f_holo e_cball
    by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros)
  moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)"
    unfolding f'_def using f_holo
    apply (intro base_residue[OF ‹open s› ‹z∈s› ‹e>0› _ e_cball,folded c_def])
    by (auto intro!:holomorphic_intros)
  ultimately have "c * f z = c * residue f' z"
    using has_contour_integral_unique by blast
  thus ?thesis unfolding c_def f'_def  by auto
qed

lemma residue_simple':
  assumes s: "open s" "z ∈ s" and holo: "f holomorphic_on (s - {z})" 
      and lim: "((λw. f w * (w - z)) ⤏ c) (at z)"
  shows   "residue f z = c"
proof -
  define g where "g = (λw. if w = z then c else f w * (w - z))"
  from holo have "(λw. f w * (w - z)) holomorphic_on (s - {z})" (is "?P")
    by (force intro: holomorphic_intros)
  also have "?P ⟷ g holomorphic_on (s - {z})"
    by (intro holomorphic_cong refl) (simp_all add: g_def)
  finally have *: "g holomorphic_on (s - {z})" .

  note lim
  also have "(λw. f w * (w - z)) ─z→ c ⟷ g ─z→ g z"
    by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter)
  finally have **: "g ─z→ g z" .

  have g_holo: "g holomorphic_on s"
    by (rule no_isolated_singularity'[where k = "{z}"])
       (insert assms * **, simp_all add: at_within_open_NO_MATCH)
  from s and this have "residue (λw. g w / (w - z)) z = g z"
    by (rule residue_simple)
  also have "∀F za in at z. g za / (za - z) = f za"
    unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def)
  hence "residue (λw. g w / (w - z)) z = residue f z"
    by (intro residue_cong refl)
  finally show ?thesis
    by (simp add: g_def)
qed




subsubsection ‹Cauchy's residue theorem›

lemma get_integrable_path:
  assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a∈s-pts" "b∈s-pts"
  obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b"
    "path_image g ⊆ s-pts" "f contour_integrable_on g" using assms
proof (induct arbitrary:s thesis a rule:finite_induct[OF ‹finite pts›])
  case 1
  obtain g where "valid_path g" "path_image g ⊆ s" "pathstart g = a" "pathfinish g = b"
    using connected_open_polynomial_connected[OF ‹open s›,of a b ] ‹connected (s - {})›
      valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto
  moreover have "f contour_integrable_on g"
    using contour_integrable_holomorphic_simple[OF _ ‹open s› ‹valid_path g› ‹path_image g ⊆ s›,of f]
      ‹f holomorphic_on s - {}›
    by auto
  ultimately show ?case using "1"(1)[of g] by auto
next
  case idt:(2 p pts)
  obtain e where "e>0" and e:"∀w∈ball a e. w ∈ s ∧ (w ≠ a ⟶ w ∉ insert p pts)"
    using finite_ball_avoid[OF ‹open s› ‹finite (insert p pts)›, of a]
      ‹a ∈ s - insert p pts›
    by auto
  define a' where "a' ≡ a+e/2"
  have "a'∈s-{p} -pts"  using e[rule_format,of "a+e/2"] ‹e>0›
    by (auto simp add:dist_complex_def a'_def)
  then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b"
    "path_image g' ⊆ s - {p} - pts" "f contour_integrable_on g'"
    using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1)
    by (metis Diff_insert2 open_delete)
  define g where "g ≡ linepath a a' +++ g'"
  have "valid_path g" unfolding g_def by (auto intro: valid_path_join)
  moreover have "pathstart g = a" and  "pathfinish g = b" unfolding g_def by auto
  moreover have "path_image g ⊆ s - insert p pts" unfolding g_def
    proof (rule subset_path_image_join)
      have "closed_segment a a' ⊆ ball a e" using ‹e>0›
        by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
      then show "path_image (linepath a a') ⊆ s - insert p pts" using e idt(9)
        by auto
    next
      show "path_image g' ⊆ s - insert p pts" using g'(4) by blast
    qed
  moreover have "f contour_integrable_on g"
    proof -
      have "closed_segment a a' ⊆ ball a e" using ‹e>0›
        by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
      then have "continuous_on (closed_segment a a') f"
        using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)]
        apply (elim continuous_on_subset)
        by auto
      then have "f contour_integrable_on linepath a a'"
        using contour_integrable_continuous_linepath by auto
      then show ?thesis unfolding g_def
        apply (rule contour_integrable_joinI)
        by (auto simp add: ‹e>0›)
    qed
  ultimately show ?case using idt.prems(1)[of g] by auto
qed

lemma Cauchy_theorem_aux:
  assumes "open s" "connected (s-pts)" "finite pts" "pts ⊆ s" "f holomorphic_on s-pts"
          "valid_path g" "pathfinish g = pathstart g" "path_image g ⊆ s-pts"
          "∀z. (z ∉ s) ⟶ winding_number g z  = 0"
          "∀p∈s. h p>0 ∧ (∀w∈cball p (h p). w∈s ∧ (w≠p ⟶ w ∉ pts))"
  shows "contour_integral g f = (∑p∈pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
    using assms
proof (induct arbitrary:s g rule:finite_induct[OF ‹finite pts›])
  case 1
  then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique)
next
  case (2 p pts)
  note fin[simp] = ‹finite (insert p pts)›
    and connected = ‹connected (s - insert p pts)›
    and valid[simp] = ‹valid_path g›
    and g_loop[simp] = ‹pathfinish g = pathstart g›
    and holo[simp]= ‹f holomorphic_on s - insert p pts›
    and path_img = ‹path_image g ⊆ s - insert p pts›
    and winding = ‹∀z. z ∉ s ⟶ winding_number g z = 0›
    and h = ‹∀pa∈s. 0 < h pa ∧ (∀w∈cball pa (h pa). w ∈ s ∧ (w ≠ pa ⟶ w ∉ insert p pts))›
  have "h p>0" and "p∈s"
    and h_p: "∀w∈cball p (h p). w ∈ s ∧ (w ≠ p ⟶ w ∉ insert p pts)"
    using h ‹insert p pts ⊆ s› by auto
  obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p"
      "path_image pg ⊆ s-insert p pts" "f contour_integrable_on pg"
    proof -
      have "p + h p∈cball p (h p)" using h[rule_format,of p]
        by (simp add: ‹p ∈ s› dist_norm)
      then have "p + h p ∈ s - insert p pts" using h[rule_format,of p] ‹insert p pts ⊆ s›
        by fastforce
      moreover have "pathstart g ∈ s - insert p pts " using path_img by auto
      ultimately show ?thesis
        using get_integrable_path[OF ‹open s› connected fin holo,of "pathstart g" "p+h p"] that
        by blast
    qed
  obtain n::int where "n=winding_number g p"
    using integer_winding_number[OF _ g_loop,of p] valid path_img
    by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path)
  define p_circ where "p_circ ≡ circlepath p (h p)"
  define p_circ_pt where "p_circ_pt ≡ linepath (p+h p) (p+h p)"
  define n_circ where "n_circ ≡ λn. (op +++ p_circ ^^ n) p_circ_pt"
  define cp where "cp ≡ if n≥0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))"
  have n_circ:"valid_path (n_circ k)"
      "winding_number (n_circ k) p = k"
      "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p"
      "path_image (n_circ k) =  (if k=0 then {p + h p} else sphere p (h p))"
      "p ∉ path_image (n_circ k)"
      "⋀p'. p'∉s - pts ⟹ winding_number (n_circ k) p'=0 ∧ p'∉path_image (n_circ k)"
      "f contour_integrable_on (n_circ k)"
      "contour_integral (n_circ k) f = k *  contour_integral p_circ f"
      for k
    proof (induct k)
      case 0
      show "valid_path (n_circ 0)"
        and "path_image (n_circ 0) =  (if 0=0 then {p + h p} else sphere p (h p))"
        and "winding_number (n_circ 0) p = of_nat 0"
        and "pathstart (n_circ 0) = p + h p"
        and "pathfinish (n_circ 0) = p + h p"
        and "p ∉ path_image (n_circ 0)"
        unfolding n_circ_def p_circ_pt_def using ‹h p > 0›
        by (auto simp add: dist_norm)
      show "winding_number (n_circ 0) p'=0 ∧ p'∉path_image (n_circ 0)" when "p'∉s- pts" for p'
        unfolding n_circ_def p_circ_pt_def
        apply (auto intro!:winding_number_trivial)
        by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+
      show "f contour_integrable_on (n_circ 0)"
        unfolding n_circ_def p_circ_pt_def
        by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing)
      show "contour_integral (n_circ 0) f = of_nat 0  *  contour_integral p_circ f"
        unfolding n_circ_def p_circ_pt_def by auto
    next
      case (Suc k)
      have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto
      have pcirc:"p ∉ path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)"
        using Suc(3) unfolding p_circ_def using ‹h p > 0› by (auto simp add: p_circ_def)
      have pcirc_image:"path_image p_circ ⊆ s - insert p pts"
        proof -
          have "path_image p_circ ⊆ cball p (h p)" using ‹0 < h p› p_circ_def by auto
          then show ?thesis using h_p pcirc(1) by auto
        qed
      have pcirc_integrable:"f contour_integrable_on p_circ"
        by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def]
          contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on
          holomorphic_on_subset[OF holo])
      show "valid_path (n_circ (Suc k))"
        using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto
      show "path_image (n_circ (Suc k))
          = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))"
        proof -
          have "path_image p_circ = sphere p (h p)"
            unfolding p_circ_def using ‹0 < h p› by auto
          then show ?thesis unfolding n_Suc  using Suc.hyps(5)  ‹h p>0›
            by (auto simp add:  path_image_join[OF pcirc(3)]  dist_norm)
        qed
      then show "p ∉ path_image (n_circ (Suc k))" using ‹h p>0› by auto
      show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)"
        proof -
          have "winding_number p_circ p = 1"
            by (simp add: ‹h p > 0› p_circ_def winding_number_circlepath_centre)
          moreover have "p ∉ path_image (n_circ k)" using Suc(5) ‹h p>0› by auto
          then have "winding_number (p_circ +++ n_circ k) p
              = winding_number p_circ p + winding_number (n_circ k) p"
            using  valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc
            apply (intro winding_number_join)
            by auto
          ultimately show ?thesis using Suc(2) unfolding n_circ_def
            by auto
        qed
      show "pathstart (n_circ (Suc k)) = p + h p"
        by (simp add: n_circ_def p_circ_def)
      show "pathfinish (n_circ (Suc k)) = p + h p"
        using Suc(4) unfolding n_circ_def by auto
      show "winding_number (n_circ (Suc k)) p'=0 ∧  p'∉path_image (n_circ (Suc k))" when "p'∉s-pts" for p'
        proof -
          have " p' ∉ path_image p_circ" using ‹p ∈ s› h p_circ_def that using pcirc_image by blast
          moreover have "p' ∉ path_image (n_circ k)"
            using Suc.hyps(7) that by blast
          moreover have "winding_number p_circ p' = 0"
            proof -
              have "path_image p_circ ⊆ cball p (h p)"
                using h unfolding p_circ_def using ‹p ∈ s› by fastforce
              moreover have "p'∉cball p (h p)" using ‹p ∈ s› h that "2.hyps"(2) by fastforce
              ultimately show ?thesis unfolding p_circ_def
                apply (intro winding_number_zero_outside)
                by auto
            qed
          ultimately show ?thesis
            unfolding n_Suc
            apply (subst winding_number_join)
            by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that])
        qed
      show "f contour_integrable_on (n_circ (Suc k))"
        unfolding n_Suc
        by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)])
      show "contour_integral (n_circ (Suc k)) f = (Suc k) *  contour_integral p_circ f"
        unfolding n_Suc
        by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]
          Suc(9) algebra_simps)
    qed
  have cp[simp]:"pathstart cp = p + h p"  "pathfinish cp = p + h p"
         "valid_path cp" "path_image cp ⊆ s - insert p pts"
         "winding_number cp p = - n"
         "⋀p'. p'∉s - pts ⟹ winding_number cp p'=0 ∧ p' ∉ path_image cp"
         "f contour_integrable_on cp"
         "contour_integral cp f = - n * contour_integral p_circ f"
    proof -
      show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp"
        using n_circ unfolding cp_def by auto
    next
      have "sphere p (h p) ⊆  s - insert p pts"
        using h[rule_format,of p] ‹insert p pts ⊆ s› by force
      moreover  have "p + complex_of_real (h p) ∈ s - insert p pts"
        using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE)
      ultimately show "path_image cp ⊆  s - insert p pts" unfolding cp_def
        using n_circ(5)  by auto
    next
      show "winding_number cp p = - n"
        unfolding cp_def using winding_number_reversepath n_circ ‹h p>0›
        by (auto simp: valid_path_imp_path)
    next
      show "winding_number cp p'=0 ∧ p' ∉ path_image cp" when "p'∉s - pts" for p'
        unfolding cp_def
        apply (auto)
        apply (subst winding_number_reversepath)
        by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1))
    next
      show "f contour_integrable_on cp" unfolding cp_def
        using contour_integrable_reversepath_eq n_circ(1,8) by auto
    next
      show "contour_integral cp f = - n * contour_integral p_circ f"
        unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9)
        by auto
    qed
  define g' where "g' ≡ g +++ pg +++ cp +++ (reversepath pg)"
  have "contour_integral g' f = (∑p∈pts. winding_number g' p * contour_integral (circlepath p (h p)) f)"
    proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ ‹finite pts› ])
      show "connected (s - {p} - pts)" using connected by (metis Diff_insert2)
      show "open (s - {p})" using ‹open s› by auto
      show " pts ⊆ s - {p}" using ‹insert p pts ⊆ s› ‹ p ∉ pts›  by blast
      show "f holomorphic_on s - {p} - pts" using holo ‹p ∉ pts› by (metis Diff_insert2)
      show "valid_path g'"
        unfolding g'_def cp_def using n_circ valid pg g_loop
        by (auto intro!:valid_path_join )
      show "pathfinish g' = pathstart g'"
        unfolding g'_def cp_def using pg(2) by simp
      show "path_image g' ⊆ s - {p} - pts"
        proof -
          define s' where "s' ≡ s - {p} - pts"
          have s':"s' = s-insert p pts " unfolding s'_def by auto
          then show ?thesis using path_img pg(4) cp(4)
            unfolding g'_def
            apply (fold s'_def s')
            apply (intro subset_path_image_join)
            by auto
        qed
      note path_join_imp[simp]
      show "∀z. z ∉ s - {p} ⟶ winding_number g' z = 0"
        proof clarify
          fix z assume z:"z∉s - {p}"
          have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z
              + winding_number (pg +++ cp +++ (reversepath pg)) z"
            proof (rule winding_number_join)
              show "path g" using ‹valid_path g› by (simp add: valid_path_imp_path)
              show "z ∉ path_image g" using z path_img by auto
              show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp
                by (simp add: valid_path_imp_path)
            next
              have "path_image (pg +++ cp +++ reversepath pg) ⊆ s - insert p pts"
                using pg(4) cp(4) by (auto simp:subset_path_image_join)
              then show "z ∉ path_image (pg +++ cp +++ reversepath pg)" using z by auto
            next
              show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto
            qed
          also have "... = winding_number g z + (winding_number pg z
              + winding_number (cp +++ (reversepath pg)) z)"
            proof (subst add_left_cancel,rule winding_number_join)
              show "path pg" and "path (cp +++ reversepath pg)"
               and "pathfinish pg = pathstart (cp +++ reversepath pg)"
                by (auto simp add: valid_path_imp_path)
              show "z ∉ path_image pg" using pg(4) z by blast
              show "z ∉ path_image (cp +++ reversepath pg)" using z
                by (metis Diff_iff ‹z ∉ path_image pg› contra_subsetD cp(4) insertI1
                  not_in_path_image_join path_image_reversepath singletonD)
            qed
          also have "... = winding_number g z + (winding_number pg z
              + (winding_number cp z + winding_number (reversepath pg) z))"
            apply (auto intro!:winding_number_join simp: valid_path_imp_path)
            apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z)
            by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z)
          also have "... = winding_number g z + winding_number cp z"
            apply (subst winding_number_reversepath)
            apply (auto simp: valid_path_imp_path)
            by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z)
          finally have "winding_number g' z = winding_number g z + winding_number cp z"
            unfolding g'_def .
          moreover have "winding_number g z + winding_number cp z = 0"
            using winding z ‹n=winding_number g p› by auto
          ultimately show "winding_number g' z = 0" unfolding g'_def by auto
        qed
      show "∀pa∈s - {p}. 0 < h pa ∧ (∀w∈cball pa (h pa). w ∈ s - {p} ∧ (w ≠ pa ⟶ w ∉ pts))"
        using h by fastforce
    qed
  moreover have "contour_integral g' f = contour_integral g f
      - winding_number g p * contour_integral p_circ f"
    proof -
      have "contour_integral g' f =  contour_integral g f
        + contour_integral (pg +++ cp +++ reversepath pg) f"
        unfolding g'_def
        apply (subst contour_integral_join)
        by (auto simp add:open_Diff[OF ‹open s›,OF finite_imp_closed[OF fin]]
          intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img]
          contour_integrable_reversepath)
      also have "... = contour_integral g f + contour_integral pg f
          + contour_integral (cp +++ reversepath pg) f"
        apply (subst contour_integral_join)
        by (auto simp add:contour_integrable_reversepath)
      also have "... = contour_integral g f + contour_integral pg f
          + contour_integral cp f + contour_integral (reversepath pg) f"
        apply (subst contour_integral_join)
        by (auto simp add:contour_integrable_reversepath)
      also have "... = contour_integral g f + contour_integral cp f"
        using contour_integral_reversepath
        by (auto simp add:contour_integrable_reversepath)
      also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f"
        using ‹n=winding_number g p› by auto
      finally show ?thesis .
    qed
  moreover have "winding_number g' p' = winding_number g p'" when "p'∈pts" for p'
    proof -
      have [simp]: "p' ∉ path_image g" "p' ∉ path_image pg" "p'∉path_image cp"
        using "2.prems"(8) that
        apply blast
        apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that)
        by (meson DiffD2 cp(4) set_rev_mp subset_insertI that)
      have "winding_number g' p' = winding_number g p'
          + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def
        apply (subst winding_number_join)
        apply (simp_all add: valid_path_imp_path)
        apply (intro not_in_path_image_join)
        by auto
      also have "... = winding_number g p' + winding_number pg p'
          + winding_number (cp +++ reversepath pg) p'"
        apply (subst winding_number_join)
        apply (simp_all add: valid_path_imp_path)
        apply (intro not_in_path_image_join)
        by auto
      also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p'
          + winding_number (reversepath pg) p'"
        apply (subst winding_number_join)
        by (simp_all add: valid_path_imp_path)
      also have "... = winding_number g p' + winding_number cp p'"
        apply (subst winding_number_reversepath)
        by (simp_all add: valid_path_imp_path)
      also have "... = winding_number g p'" using that by auto
      finally show ?thesis .
    qed
  ultimately show ?case unfolding p_circ_def
    apply (subst (asm) sum.cong[OF refl,
        of pts _ "λp. winding_number g p * contour_integral (circlepath p (h p)) f"])
    by (auto simp add:sum.insert[OF ‹finite pts› ‹p∉pts›] algebra_simps)
qed

lemma Cauchy_theorem_singularities:
  assumes "open s" "connected s" "finite pts" and
          holo:"f holomorphic_on s-pts" and
          "valid_path g" and
          loop:"pathfinish g = pathstart g" and
          "path_image g ⊆ s-pts" and
          homo:"∀z. (z ∉ s) ⟶ winding_number g z  = 0" and
          avoid:"∀p∈s. h p>0 ∧ (∀w∈cball p (h p). w∈s ∧ (w≠p ⟶ w ∉ pts))"
  shows "contour_integral g f = (∑p∈pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
    (is "?L=?R")
proof -
  define circ where "circ ≡ λp. winding_number g p * contour_integral (circlepath p (h p)) f"
  define pts1 where "pts1 ≡ pts ∩ s"
  define pts2 where "pts2 ≡ pts - pts1"
  have "pts=pts1 ∪ pts2" "pts1 ∩ pts2 = {}" "pts2 ∩ s={}" "pts1⊆s"
    unfolding pts1_def pts2_def by auto
  have "contour_integral g f =  (∑p∈pts1. circ p)" unfolding circ_def
    proof (rule Cauchy_theorem_aux[OF ‹open s› _ _ ‹pts1⊆s› _ ‹valid_path g› loop _ homo])
      have "finite pts1" unfolding pts1_def using ‹finite pts› by auto
      then show "connected (s - pts1)"
        using ‹open s› ‹connected s› connected_open_delete_finite[of s] by auto
    next
      show "finite pts1" using ‹pts = pts1 ∪ pts2› assms(3) by auto
      show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def)
      show "path_image g ⊆ s - pts1" using assms(7) pts1_def by auto
      show "∀p∈s. 0 < h p ∧ (∀w∈cball p (h p). w ∈ s ∧ (w ≠ p ⟶ w ∉ pts1))"
        by (simp add: avoid pts1_def)
    qed
  moreover have "sum circ pts2=0"
    proof -
      have "winding_number g p=0" when "p∈pts2" for p
        using  ‹pts2 ∩ s={}› that homo[rule_format,of p] by auto
      thus ?thesis unfolding circ_def
        apply (intro sum.neutral)
        by auto
    qed
  moreover have "?R=sum circ pts1 + sum circ pts2"
    unfolding circ_def
    using sum.union_disjoint[OF _ _ ‹pts1 ∩ pts2 = {}›] ‹finite pts› ‹pts=pts1 ∪ pts2›
    by blast
  ultimately show ?thesis
    apply (fold circ_def)
    by auto
qed

lemma Residue_theorem:
  fixes s pts::"complex set" and f::"complex ⇒ complex"
    and g::"real ⇒ complex"
  assumes "open s" "connected s" "finite pts" and
          holo:"f holomorphic_on s-pts" and
          "valid_path g" and
          loop:"pathfinish g = pathstart g" and
          "path_image g ⊆ s-pts" and
          homo:"∀z. (z ∉ s) ⟶ winding_number g z  = 0"
  shows "contour_integral g f = 2 * pi * 𝗂 *(∑p∈pts. winding_number g p * residue f p)"
proof -
  define c where "c ≡  2 * pi * 𝗂"
  obtain h where avoid:"∀p∈s. h p>0 ∧ (∀w∈cball p (h p). w∈s ∧ (w≠p ⟶ w ∉ pts))"
    using finite_cball_avoid[OF ‹open s› ‹finite pts›] by metis
  have "contour_integral g f
      = (∑p∈pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
    using Cauchy_theorem_singularities[OF assms avoid] .
  also have "... = (∑p∈pts.  c * winding_number g p * residue f p)"
    proof (intro sum.cong)
      show "pts = pts" by simp
    next
      fix x assume "x ∈ pts"
      show "winding_number g x * contour_integral (circlepath x (h x)) f
          = c * winding_number g x * residue f x"
        proof (cases "x∈s")
          case False
          then have "winding_number g x=0" using homo by auto
          thus ?thesis by auto
        next
          case True
          have "contour_integral (circlepath x (h x)) f = c* residue f x"
            using ‹x∈pts› ‹finite pts› avoid[rule_format,OF True]
            apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def])
            by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF ‹open s› finite_imp_closed])
          then show ?thesis by auto
        qed
    qed
  also have "... = c * (∑p∈pts. winding_number g p * residue f p)"
    by (simp add: sum_distrib_left algebra_simps)
  finally show ?thesis unfolding c_def .
qed

subsection ‹The argument principle›

definition is_pole :: "('a::topological_space ⇒ 'b::real_normed_vector) ⇒ 'a ⇒ bool" where
  "is_pole f a =  (LIM x (at a). f x :> at_infinity)"

lemma is_pole_tendsto:
  fixes f::"('a::topological_space ⇒ 'b::real_normed_div_algebra)"
  shows "is_pole f x ⟹ ((inverse o f) ⤏ 0) (at x)"
unfolding is_pole_def
by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at)

lemma is_pole_inverse_holomorphic:
  assumes "open s"
    and f_holo:"f holomorphic_on (s-{z})"
    and pole:"is_pole f z"
    and non_z:"∀x∈s-{z}. f x≠0"
  shows "(λx. if x=z then 0 else inverse (f x)) holomorphic_on s"
proof -
  define g where "g ≡ λx. if x=z then 0 else inverse (f x)"
  have "isCont g z" unfolding isCont_def  using is_pole_tendsto[OF pole]
    apply (subst Lim_cong_at[where b=z and y=0 and g="inverse ∘ f"])
    by (simp_all add:g_def)
  moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto
  hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def
    by (auto elim!:continuous_on_inverse simp add:non_z)
  hence "continuous_on (s-{z}) g" unfolding g_def
    apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"])
    by auto
  ultimately have "continuous_on s g" using open_delete[OF ‹open s›] ‹open s›
    by (auto simp add:continuous_on_eq_continuous_at)
  moreover have "(inverse o f) holomorphic_on (s-{z})"
    unfolding comp_def using f_holo
    by (auto elim!:holomorphic_on_inverse simp add:non_z)
  hence "g holomorphic_on (s-{z})"
    apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"])
    by (auto simp add:g_def)
  ultimately show ?thesis unfolding g_def using ‹open s›
    by (auto elim!: no_isolated_singularity)
qed


(*order of the zero of f at z*)
definition zorder::"(complex ⇒ complex) ⇒ complex ⇒ nat" where
  "zorder f z = (THE n. n>0 ∧ (∃h r. r>0 ∧ h holomorphic_on cball z r
                    ∧ (∀w∈cball z r. f w =  h w * (w-z)^n ∧ h w ≠0)))"

definition zer_poly::"[complex ⇒ complex,complex]⇒complex ⇒ complex" where
  "zer_poly f z = (SOME h. ∃r . r>0 ∧ h holomorphic_on cball z r
                    ∧ (∀w∈cball z r. f w =  h w * (w-z)^(zorder f z) ∧ h w ≠0))"

(*order of the pole of f at z*)
definition porder::"(complex ⇒ complex) ⇒ complex ⇒ nat" where
  "porder f z = (let f'=(λx. if x=z then 0 else inverse (f x)) in zorder f' z)"

definition pol_poly::"[complex ⇒ complex,complex]⇒complex ⇒ complex" where
  "pol_poly f z = (let f'=(λ x. if x=z then 0 else inverse (f x))
      in inverse o zer_poly f' z)"


lemma holomorphic_factor_zero_unique:
  fixes f::"complex ⇒ complex" and z::complex and r::real
  assumes "r>0"
    and asm:"∀w∈ball z r. f w = (w-z)^n * g w ∧ g w≠0 ∧ f w = (w - z)^m * h w ∧ h w≠0"
    and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r"
  shows "n=m"
proof -
  have "n>m ⟹ False"
    proof -
      assume "n>m"
      have "(h ⤏ 0) (at z within ball z r)"
        proof (rule Lim_transform_within[OF _ ‹r>0›, where f="λw. (w - z) ^ (n - m) * g w"])
          have "∀w∈ball z r. w≠z ⟶ h w = (w-z)^(n-m) * g w" using ‹n>m› asm
            by (auto simp add:field_simps power_diff)
          then show "⟦x' ∈ ball z r; 0 < dist x' z;dist x' z < r⟧
            ⟹ (x' - z) ^ (n - m) * g x' = h x'" for x' by auto
        next
          define F where "F ≡ at z within ball z r"
          define f' where "f' ≡ λx. (x - z) ^ (n-m)"
          have "f' z=0" using ‹n>m› unfolding f'_def by auto
          moreover have "continuous F f'" unfolding f'_def F_def
            by (intro continuous_intros)
          ultimately have "(f' ⤏ 0) F" unfolding F_def
            by (simp add: continuous_within)
          moreover have "(g ⤏ g z) F"
            using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] ‹r>0›
            unfolding F_def by auto
          ultimately show " ((λw. f' w * g w) ⤏ 0) F" using tendsto_mult by fastforce
        qed
      moreover have "(h ⤏ h z) (at z within ball z r)"
        using holomorphic_on_imp_continuous_on[OF h_holo]
        by (auto simp add:continuous_on_def ‹r>0›)
      moreover have "at z within ball z r ≠ bot" using ‹r>0›
        by (auto simp add:trivial_limit_within islimpt_ball)
      ultimately have "h z=0" by (auto intro: tendsto_unique)
      thus False using asm ‹r>0› by auto
    qed
  moreover have "m>n ⟹ False"
    proof -
      assume "m>n"
      have "(g ⤏ 0) (at z within ball z r)"
        proof (rule Lim_transform_within[OF _ ‹r>0›, where f="λw. (w - z) ^ (m - n) * h w"])
          have "∀w∈ball z r. w≠z ⟶ g w = (w-z)^(m-n) * h w" using ‹m>n› asm
            by (auto simp add:field_simps power_diff)
          then show "⟦x' ∈ ball z r; 0 < dist x' z;dist x' z < r⟧
            ⟹ (x' - z) ^ (m - n) * h x' = g x'" for x' by auto
        next
          define F where "F ≡ at z within ball z r"
          define f' where "f' ≡λx. (x - z) ^ (m-n)"
          have "f' z=0" using ‹m>n› unfolding f'_def by auto
          moreover have "continuous F f'" unfolding f'_def F_def
            by (intro continuous_intros)
          ultimately have "(f' ⤏ 0) F" unfolding F_def
            by (simp add: continuous_within)
          moreover have "(h ⤏ h z) F"
            using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] ‹r>0›
            unfolding F_def by auto
          ultimately show " ((λw. f' w * h w) ⤏ 0) F" using tendsto_mult by fastforce
        qed
      moreover have "(g ⤏ g z) (at z within ball z r)"
        using holomorphic_on_imp_continuous_on[OF g_holo]
        by (auto simp add:continuous_on_def ‹r>0›)
      moreover have "at z within ball z r ≠ bot" using ‹r>0›
        by (auto simp add:trivial_limit_within islimpt_ball)
      ultimately have "g z=0" by (auto intro: tendsto_unique)
      thus False using asm ‹r>0› by auto
    qed
  ultimately show "n=m" by fastforce
qed


lemma holomorphic_factor_zero_Ex1:
  assumes "open s" "connected s" "z ∈ s" and
        holo:"f holomorphic_on s"
        and "f z = 0" and "∃w∈s. f w ≠ 0"
  shows "∃!n. ∃g r. 0 < n ∧ 0 < r ∧
                g holomorphic_on cball z r
                ∧ (∀w∈cball z r. f w = (w-z)^n * g w ∧ g w≠0)"
proof (rule ex_ex1I)
  obtain g r n where "0 < n" "0 < r" "ball z r ⊆ s" and
          g:"g holomorphic_on ball z r"
          "⋀w. w ∈ ball z r ⟹ f w = (w - z) ^ n * g w"
          "⋀w. w ∈ ball z r ⟹ g w ≠ 0"
    using holomorphic_factor_zero_nonconstant[OF holo ‹open s› ‹connected s› ‹z∈s› ‹f z=0›]
    by (metis assms(3) assms(5) assms(6))
  define r' where "r' ≡ r/2"
  have "cball z r' ⊆ ball z r" unfolding r'_def by (simp add: ‹0 < r› cball_subset_ball_iff)
  hence "cball z r' ⊆ s" "g holomorphic_on cball z r'"
      "(∀w∈cball z r'. f w = (w - z) ^ n * g w ∧ g w ≠ 0)"
    using g ‹ball z r ⊆ s› by auto
  moreover have "r'>0" unfolding r'_def using ‹0<r› by auto
  ultimately show "∃n g r. 0 < n ∧ 0 < r  ∧ g holomorphic_on cball z r
          ∧ (∀w∈cball z r. f w = (w - z) ^ n * g w ∧ g w ≠ 0)"
    apply (intro exI[of _ n] exI[of _ g] exI[of _ r'])
    by (simp add:‹0 < n›)
next
  fix m n
  define fac where "fac ≡ λn g r. ∀w∈cball z r. f w = (w - z) ^ n * g w ∧ g w ≠ 0"
  assume n_asm:"∃g r1. 0 < n ∧ 0 < r1 ∧ g holomorphic_on cball z r1 ∧ fac n g r1"
     and m_asm:"∃h r2. 0 < m ∧ 0 < r2  ∧ h holomorphic_on cball z r2 ∧ fac m h r2"
  obtain g r1 where "0 < n" "0 < r1" and g_holo: "g holomorphic_on cball z r1"
    and "fac n g r1" using n_asm by auto
  obtain h r2 where "0 < m" "0 < r2" and h_holo: "h holomorphic_on cball z r2"
    and "fac m h r2" using m_asm by auto
  define r where "r ≡ min r1 r2"
  have "r>0" using ‹r1>0› ‹r2>0› unfolding r_def by auto
  moreover have "∀w∈ball z r. f w = (w-z)^n * g w ∧ g w≠0 ∧ f w = (w - z)^m * h w ∧ h w≠0"
    using ‹fac m h r2› ‹fac n g r1›   unfolding fac_def r_def
    by fastforce
  ultimately show "m=n" using g_holo h_holo
    apply (elim holomorphic_factor_zero_unique[of r z f n g m h,symmetric,rotated])
    by (auto simp add:r_def)
qed

lemma zorder_exist:
  fixes f::"complex ⇒ complex" and z::complex
  defines "n≡zorder f z" and "h≡zer_poly f z"
  assumes  "open s" "connected s" "z∈s"
    and holo: "f holomorphic_on s"
    and  "f z=0" "∃w∈s. f w≠0"
  shows "∃r. n>0 ∧ r>0 ∧ cball z r ⊆ s ∧ h holomorphic_on cball z r
    ∧ (∀w∈cball z r. f w  = h w * (w-z)^n ∧ h w ≠0) "
proof -
  define P where "P ≡ λh r n. r>0 ∧ h holomorphic_on cball z r
    ∧ (∀w∈cball z r. ( f w  = h w * (w-z)^n) ∧ h w ≠0)"
  have "(∃!n. n>0 ∧ (∃ h r. P h r n))"
    proof -
      have "∃!n. ∃h r. n>0 ∧ P h r n"
        using holomorphic_factor_zero_Ex1[OF ‹open s› ‹connected s› ‹z∈s› holo ‹f z=0›
          ‹∃w∈s. f w≠0›] unfolding P_def
        apply (subst mult.commute)
        by auto
      thus ?thesis by auto
    qed
  moreover have n:"n=(THE n. n>0 ∧ (∃h r. P h r n))"
    unfolding n_def zorder_def P_def by simp
  ultimately have "n>0 ∧ (∃h r. P h r n)"
    apply (drule_tac theI')
    by simp
  then have "n>0" and "∃h r. P h r n" by auto
  moreover have "h=(SOME h. ∃r. P h r n)"
    unfolding h_def P_def zer_poly_def[of f z,folded n_def P_def] by simp
  ultimately have "∃r. P h r n"
    apply (drule_tac someI_ex)
    by simp
  then obtain r1 where "P h r1 n" by auto
  obtain r2 where "r2>0" "cball z r2 ⊆ s"
    using assms(3) assms(5) open_contains_cball_eq by blast
  define r3 where "r3 ≡ min r1 r2"
  have "P h r3 n" using ‹P h r1 n› ‹r2>0› unfolding P_def r3_def
    by auto
  moreover have "cball z r3 ⊆ s" using ‹cball z r2 ⊆ s› unfolding r3_def by auto
  ultimately show ?thesis using ‹n>0› unfolding P_def by auto
qed

lemma zorder_eqI:
  assumes "open s" "z ∈ s" "g holomorphic_on s" "g z ≠ 0" "n > 0"
  assumes "⋀w. w ∈ s ⟹ f w = g w * (w - z) ^ n"
  shows   "zorder f z = n"
proof -
  have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact
  moreover have "open (-{0::complex})" by auto
  ultimately have "open ((g -` (-{0})) ∩ s)"
    unfolding continuous_on_open_vimage[OF ‹open s›] by blast
  moreover from assms have "z ∈ (g -` (-{0})) ∩ s" by auto
  ultimately obtain r where r: "r > 0" "cball z r ⊆ (g -` (-{0})) ∩ s"
    unfolding open_contains_cball by blast

  have "n > 0 ∧ r > 0 ∧ g holomorphic_on cball z r ∧ 
        (∀w∈cball z r. f w = (w - z) ^ n * g w ∧ g w ≠ 0)" (is "?P g r n")
    using r assms(3,5,6) by auto
  hence ex: "∃g r. ?P g r n" by blast
  have unique: "∃!n. ∃g r. ?P g r n"
  proof (rule holomorphic_factor_zero_Ex1)
    from r have "(λw. g w * (w - z) ^ n) holomorphic_on ball z r"
      by (intro holomorphic_intros holomorphic_on_subset[OF assms(3)]) auto
    also have "?this ⟷ f holomorphic_on ball z r"
      using r assms by (intro holomorphic_cong refl) (auto simp: cball_def subset_iff)
    finally show  .
  next
    let ?w = "z + of_real r / 2"
    have "?w ∈ ball z r"
      using r by (auto simp: dist_norm)
    moreover from this and r have "g ?w ≠ 0" and "?w ∈ s" 
      by (auto simp: cball_def ball_def subset_iff)
    with assms have "f ?w ≠ 0" using ‹r > 0› by auto
    ultimately show "∃w∈ball z r. f w ≠ 0" by blast
  qed (insert assms r, auto)
  from unique and ex have "(THE n. ∃g r. ?P g r n) = n"
    by (rule the1_equality)
  also have "(THE n. ∃g r. ?P g r n) = zorder f z"
    by (simp add: zorder_def mult.commute)
  finally show ?thesis .
qed

lemma simple_zeroI:
  assumes "open s" "z ∈ s" "g holomorphic_on s" "g z ≠ 0"
  assumes "⋀w. w ∈ s ⟹ f w = g w * (w - z)"
  shows   "zorder f z = 1"
  using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto)

lemma higher_deriv_power:
  shows   "(deriv ^^ j) (λw. (w - z) ^ n) w = 
             pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)"
proof (induction j arbitrary: w)
  case 0
  thus ?case by auto
next
  case (Suc j w)
  have "(deriv ^^ Suc j) (λw. (w - z) ^ n) w = deriv ((deriv ^^ j) (λw. (w - z) ^ n)) w"
    by simp
  also have "(deriv ^^ j) (λw. (w - z) ^ n) = 
               (λw. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))"
    using Suc by (intro Suc.IH ext)
  also {
    have "(… has_field_derivative of_nat (n - j) *
               pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)"
      using Suc.prems by (auto intro!: derivative_eq_intros)
    also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = 
                 pochhammer (of_nat (Suc n - Suc j)) (Suc j)"
      by (cases "Suc j ≤ n", subst pochhammer_rec) 
         (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left)
    finally have "deriv (λw. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w =
                    … * (w - z) ^ (n - Suc j)"
      by (rule DERIV_imp_deriv)
  }
  finally show ?case .
qed

lemma zorder_eqI':
  assumes "open s" "connected s" "z ∈ s" "f holomorphic_on s"
  assumes zero: "⋀i. i < n' ⟹ (deriv ^^ i) f z = 0"
  assumes nz: "(deriv ^^ n') f z ≠ 0" and n: "n' > 0"
  shows   "zorder f z = n'"
proof -
  {
    assume *: "⋀w. w ∈ s ⟹ f w = 0"
    hence "eventually (λu. u ∈ s) (nhds z)"
      using assms by (intro eventually_nhds_in_open) auto
    hence "eventually (λu. f u = 0) (nhds z)"
      by eventually_elim (simp_all add: *)
    hence "(deriv ^^ n') f z = (deriv ^^ n') (λ_. 0) z"
      by (intro higher_deriv_cong_ev) auto
    also have "(deriv ^^ n') (λ_. 0) z = 0"
      by (induction n') simp_all
    finally have False using nz by contradiction
  }
  hence nz': "∃w∈s. f w ≠ 0" by blast

  from zero[of 0] and n have [simp]: "f z = 0" by simp

  define n g where "n = zorder f z" and "g = zer_poly f z"
  from zorder_exist[OF assms(1-4) ‹f z = 0› nz']
    obtain r where r: "n > 0" "r > 0" "cball z r ⊆ s" "g holomorphic_on cball z r"
                      "∀w∈cball z r. f w = g w * (w - z) ^ n ∧ g w ≠ 0"
    unfolding n_def g_def by blast

  define A where "A = (λi. of_nat (i choose n) * fact n * (deriv ^^ (i - n)) g z)"
  {
    fix i :: nat
    have "eventually (λw. w ∈ ball z r) (nhds z)"
      using r by (intro eventually_nhds_in_open) auto
    hence "eventually (λw. f w = (w - z) ^ n * g w) (nhds z)"
      by eventually_elim (use r in auto)
    hence "(deriv ^^ i) f z = (deriv ^^ i) (λw. (w - z) ^ n * g w) z"
      by (intro higher_deriv_cong_ev) auto
    also have "… = (∑j=0..i. of_nat (i choose j) *
                       (deriv ^^ j) (λw. (w - z) ^ n) z * (deriv ^^ (i - j)) g z)"
      using r by (intro higher_deriv_mult[of _ "ball z r"]) (auto intro!: holomorphic_intros)
    also have "… = (∑j=0..i. if j = n then of_nat (i choose n) * fact n * (deriv ^^ (i - n)) g z 
                                 else 0)"
    proof (intro sum.cong refl, goal_cases)
      case (1 j)
      have "(deriv ^^ j) (λw. (w - z) ^ n) z = 
              pochhammer (of_nat (Suc n - j)) j * 0 ^ (n - j)"
        by (subst higher_deriv_power) auto
      also have "… = (if j = n then fact j else 0)"
        by (auto simp: not_less pochhammer_0_left pochhammer_fact)
      also have "of_nat (i choose j) * … * (deriv ^^ (i - j)) g z = 
                   (if j = n then of_nat (i choose n) * fact n * (deriv ^^ (i - n)) g z else 0)"
        by simp
      finally show ?case .
    qed
    also have "… = (if i ≥ n then A i else 0)"
      by (auto simp: A_def)
    finally have "(deriv ^^ i) f z = …" .
  } note * = this

  from *[of n] and r have "(deriv ^^ n) f z ≠ 0"
    by (simp add: A_def)
  with zero[of n] have "n ≥ n'" by (cases "n ≥ n'") auto
  with nz show "n = n'"
    by (auto simp: * split: if_splits)
qed

lemma simple_zeroI':
  assumes "open s" "connected s" "z ∈ s" 
  assumes "⋀z. z ∈ s ⟹ (f has_field_derivative f' z) (at z)"
  assumes "f z = 0" "f' z ≠ 0"
  shows   "zorder f z = 1"
proof -
  have "deriv f z = f' z" if "z ∈ s" for z
    using that by (intro DERIV_imp_deriv assms) auto
  moreover from assms have "f holomorphic_on s"
    by (subst holomorphic_on_open) auto
  ultimately show ?thesis using assms
    by (intro zorder_eqI'[of s]) auto
qed

lemma porder_exist:
  fixes f::"complex ⇒ complex" and z::complex
  defines "n ≡ porder f z" and "h ≡ pol_poly f z"
  assumes "open s" "z ∈ s"
    and holo:"f holomorphic_on s-{z}"
    and "is_pole f z"
  shows "∃r. n>0 ∧ r>0 ∧ cball z r ⊆ s ∧ h holomorphic_on cball z r
    ∧ (∀w∈cball z r. (w≠z ⟶ f w  = h w / (w-z)^n) ∧ h w ≠0)"
proof -
  obtain e where "e>0" and e_ball:"ball z e ⊆ s"and e_def: "∀x∈ball z e-{z}. f x≠0"
    proof -
      have "∀F z in at z. f z ≠ 0"
        using ‹is_pole f z› filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def
        by auto
      then obtain e1 where "e1>0" and e1_def: "∀x. x ≠ z ∧ dist x z < e1 ⟶ f x ≠ 0"
        using eventually_at[of "λx. f x≠0" z,simplified] by auto
      obtain e2 where "e2>0" and "ball z e2 ⊆s" using ‹open s› ‹z∈s› openE by auto
      define e where "e ≡ min e1 e2"
      have "e>0" using ‹e1>0› ‹e2>0› unfolding e_def by auto
      moreover have "ball z e ⊆ s" unfolding e_def using ‹ball z e2 ⊆ s› by auto
      moreover have "∀x∈ball z e-{z}. f x≠0" using e1_def ‹e1>0› ‹e2>0› unfolding e_def
        by (simp add: DiffD1 DiffD2 dist_commute singletonI)
      ultimately show ?thesis using that by auto
    qed
  define g where "g ≡ λx. if x=z then 0 else inverse (f x)"
  define zo where "zo ≡ zorder g z"
  define zp where "zp ≡ zer_poly g z"
  have "∃w∈ball z e. g w ≠ 0"
    proof -
      obtain w where w:"w∈ball z e-{z}" using ‹0 < e›
        by (metis open_ball all_not_in_conv centre_in_ball insert_Diff_single
          insert_absorb not_open_singleton)
      hence "w≠z" "f w≠0" using e_def[rule_format,of w] mem_ball
        by (auto simp add:dist_commute)
      then show ?thesis unfolding g_def using w by auto
    qed
  moreover have "g holomorphic_on ball z e"
    apply (intro is_pole_inverse_holomorphic[of "ball z e",OF _ _ ‹is_pole f z› e_def,folded g_def])
    using holo e_ball by auto
  moreover have "g z=0" unfolding g_def by auto
  ultimately obtain r where "0 < zo" "0 < r" "cball z r ⊆ ball z e"
      and zp_holo: "zp holomorphic_on cball z r" and
      zp_fac: "∀w∈cball z r. g w = zp w * (w - z) ^ zo ∧ zp w ≠ 0"
    using zorder_exist[of "ball z e" z g,simplified,folded zo_def zp_def] ‹e>0›
    by auto
  have n:"n=zo" and h:"h=inverse o zp"
    unfolding n_def zo_def porder_def h_def zp_def pol_poly_def g_def by simp_all
  have "h holomorphic_on cball z r"
    using zp_holo zp_fac holomorphic_on_inverse  unfolding h comp_def by blast
  moreover have "∀w∈cball z r. (w≠z ⟶ f w  = h w / (w-z)^n) ∧ h w ≠0"
    using zp_fac unfolding h n comp_def g_def
    by (metis divide_inverse_commute field_class.field_inverse_zero inverse_inverse_eq
      inverse_mult_distrib mult.commute)
  moreover have "0 < n" unfolding n using ‹zo>0› by simp
  ultimately show ?thesis using ‹0 < r› ‹cball z r ⊆ ball z e› e_ball by auto
qed

lemma residue_porder:
  fixes f::"complex ⇒ complex" and z::complex
  defines "n ≡ porder f z" and "h ≡ pol_poly f z"
  assumes "open s" "z ∈ s"
    and holo:"f holomorphic_on s - {z}"
    and pole:"is_pole f z"
  shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))"
proof -
  define g where "g ≡ λx. if x=z then 0 else inverse (f x)"
  obtain r where "0 < n" "0 < r" and r_cball:"cball z r ⊆ s" and h_holo: "h holomorphic_on cball z r"
      and h_divide:"(∀w∈cball z r. (w≠z ⟶ f w = h w / (w - z) ^ n) ∧ h w ≠ 0)"
    using porder_exist[OF ‹open s› ‹z ∈ s› holo pole, folded n_def h_def] by blast
  have r_nonzero:"⋀w. w ∈ ball z r - {z} ⟹ f w ≠ 0"
    using h_divide by simp
  define c where "c ≡ 2 * pi * 𝗂"
  define der_f where "der_f ≡ ((deriv ^^ (n - 1)) h z / fact (n-1))"
  define h' where "h' ≡ λu. h u / (u - z) ^ n"
  have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)"
    unfolding h'_def
    proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1",
        folded c_def Suc_pred'[OF ‹n>0›]])
      show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp
      show "h holomorphic_on ball z r" using h_holo by auto
      show " z ∈ ball z r" using ‹r>0› by auto
    qed
  then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto
  then have "(f has_contour_integral c * der_f) (circlepath z r)"
    proof (elim has_contour_integral_eq)
      fix x assume "x ∈ path_image (circlepath z r)"
      hence "x∈cball z r - {z}" using ‹r>0› by auto
      then show "h' x = f x" using h_divide unfolding h'_def by auto
    qed
  moreover have "(f has_contour_integral c * residue f z) (circlepath z r)"
    using base_residue[OF ‹open s› ‹z∈s› ‹r>0› holo r_cball,folded c_def] .
  ultimately have "c * der_f =  c * residue f z" using has_contour_integral_unique by blast
  hence "der_f = residue f z" unfolding c_def by auto
  thus ?thesis unfolding der_f_def by auto
qed

theorem argument_principle:
  fixes f::"complex ⇒ complex" and poles s:: "complex set"
  defines "zeros≡{p. f p=0} - poles"
  assumes "open s" and
          "connected s" and
          f_holo:"f holomorphic_on s-poles" and
          h_holo:"h holomorphic_on s" and
          "valid_path g" and
          loop:"pathfinish g = pathstart g" and
          path_img:"path_image g ⊆ s - (zeros ∪ poles)" and
          homo:"∀z. (z ∉ s) ⟶ winding_number g z = 0" and
          finite:"finite (zeros ∪ poles)" and
          poles:"∀p∈poles. is_pole f p"
  shows "contour_integral g (λx. deriv f x * h x / f x) = 2 * pi * 𝗂 *
          ((∑p∈zeros. winding_number g p * h p * zorder f p)
           - (∑p∈poles. winding_number g p * h p * porder f p))"
    (is "?L=?R")
proof -
  define c where "c ≡ 2 * complex_of_real pi * 𝗂 "
  define ff where "ff ≡ (λx. deriv f x * h x / f x)"
  define cont_pole where "cont_pole ≡ λff p e. (ff has_contour_integral - c  * porder f p * h p) (circlepath p e)"
  define cont_zero where "cont_zero ≡ λff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)"
  define avoid where "avoid ≡ λp e. ∀w∈cball p e. w ∈ s ∧ (w ≠ p ⟶ w ∉ zeros ∪ poles)"
  have "∃e>0. avoid p e ∧ (p∈poles ⟶ cont_pole ff p e) ∧ (p∈zeros ⟶ cont_zero ff p e)"
      when "p∈s" for p
    proof -
      obtain e1 where "e1>0" and e1_avoid:"avoid p e1"
        using finite_cball_avoid[OF ‹open s› finite] ‹p∈s› unfolding avoid_def by auto
      have "∃e2>0. cball p e2 ⊆ ball p e1 ∧ cont_pole ff p e2"
        when "p∈poles"
        proof -
          define po where "po ≡ porder f p"
          define pp where "pp ≡ pol_poly f p"
          define f' where "f' ≡ λw. pp w / (w - p) ^ po"
          define ff' where "ff' ≡ (λx. deriv f' x * h x / f' x)"
          have "f holomorphic_on ball p e1 - {p}"
            apply (intro holomorphic_on_subset[OF f_holo])
            using e1_avoid ‹p∈poles› unfolding avoid_def by auto
          then obtain r where
              "0 < po" "r>0"
              "cball p r ⊆ ball p e1" and
              pp_holo:"pp holomorphic_on cball p r" and
              pp_po:"(∀w∈cball p r. (w≠p ⟶ f w = pp w / (w - p) ^ po) ∧ pp w ≠ 0)"
            using porder_exist[of "ball p e1" p f,simplified,OF ‹e1>0›] poles ‹p∈poles›
            unfolding po_def pp_def
            by auto
          define e2 where "e2 ≡ r/2"
          have "e2>0" using ‹r>0› unfolding e2_def by auto
          define anal where "anal ≡ λw. deriv pp w * h w / pp w"
          define prin where "prin ≡ λw. - of_nat po * h w / (w - p)"
          have "((λw.  prin w + anal w) has_contour_integral - c * po * h p) (circlepath p e2)"
            proof (rule  has_contour_integral_add[of _ _ _ _ 0,simplified])
              have "ball p r ⊆ s"
                using ‹cball p r ⊆ ball p e1› avoid_def ball_subset_cball e1_avoid by blast
              then have "cball p e2 ⊆ s"
                using ‹r>0› unfolding e2_def by auto
              then have "(λw. - of_nat po * h w) holomorphic_on cball p e2"
                using h_holo
                by (auto intro!: holomorphic_intros)
              then show "(prin has_contour_integral - c * of_nat po * h p ) (circlepath p e2)"
                using Cauchy_integral_circlepath_simple[folded c_def, of "λw. - of_nat po * h w"]
                  ‹e2>0›
                unfolding prin_def
                by (auto simp add: mult.assoc)
              have "anal holomorphic_on ball p r" unfolding anal_def
                using pp_holo h_holo pp_po ‹ball p r ⊆ s›
                by (auto intro!: holomorphic_intros)
              then show "(anal has_contour_integral 0) (circlepath p e2)"
                using e2_def ‹r>0›
                by (auto elim!: Cauchy_theorem_disc_simple)
            qed
          then have "cont_pole ff' p e2" unfolding cont_pole_def po_def
            proof (elim has_contour_integral_eq)
              fix w assume "w ∈ path_image (circlepath p e2)"
              then have "w∈ball p r" and "w≠p" unfolding e2_def using ‹r>0› by auto
              define wp where "wp ≡ w-p"
              have "wp≠0" and "pp w ≠0"
                unfolding wp_def using ‹w≠p› ‹w∈ball p r› pp_po by auto
              moreover have der_f':"deriv f' w = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po"
                proof (rule DERIV_imp_deriv)
                  define der where "der ≡ - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po"
                  have po:"po = Suc (po - Suc 0) " using ‹po>0› by auto
                  have "(pp has_field_derivative (deriv pp w)) (at w)"
                    using DERIV_deriv_iff_has_field_derivative pp_holo ‹w≠p›
                      by (meson open_ball ‹w ∈ ball p r› ball_subset_cball holomorphic_derivI holomorphic_on_subset)
                  then show "(f' has_field_derivative  der) (at w)"
                    using ‹w≠p› ‹po>0› unfolding der_def f'_def
                    apply (auto intro!: derivative_eq_intros simp add:field_simps)
                    apply (subst (4) po)
                    apply (subst power_Suc)
                    by (auto simp add:field_simps)
                qed
              ultimately show "prin w + anal w = ff' w"
                unfolding ff'_def prin_def anal_def
                apply simp
                apply (unfold f'_def)
                apply (fold wp_def)
                by (auto simp add:field_simps)
            qed
          then have "cont_pole ff p e2" unfolding cont_pole_def
            proof (elim has_contour_integral_eq)
              fix w assume "w ∈ path_image (circlepath p e2)"
              then have "w∈ball p r" and "w≠p" unfolding e2_def using ‹r>0› by auto
              have "deriv f' w =  deriv f w"
                proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"])
                  show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo
                    by (auto intro!: holomorphic_intros)
                next
                  have "ball p e1 - {p} ⊆ s - poles"
                    using avoid_def ball_subset_cball e1_avoid
                    by auto
                  then have "ball p r - {p} ⊆ s - poles" using ‹cball p r ⊆ ball p e1›
                    using ball_subset_cball by blast
                  then show "f holomorphic_on ball p r - {p}" using f_holo
                    by auto
                next
                  show "open (ball p r - {p})" by auto
                next
                  show "w ∈ ball p r - {p}" using ‹w∈ball p r› ‹w≠p› by auto
                next
                  fix x assume "x ∈ ball p r - {p}"
                  then show "f' x = f x"
                    using pp_po unfolding f'_def by auto
                qed
              moreover have " f' w  =  f w "
                using ‹w ∈ ball p r› ball_subset_cball subset_iff pp_po ‹w≠p›
                unfolding f'_def by auto
              ultimately show "ff' w = ff w"
                unfolding ff'_def ff_def by simp
            qed
          moreover have "cball p e2 ⊆ ball p e1"
            using ‹0 < r› ‹cball p r ⊆ ball p e1› e2_def by auto
          ultimately show ?thesis using ‹e2>0› by auto
        qed
      then obtain e2 where e2:"p∈poles ⟶ e2>0 ∧ cball p e2 ⊆ ball p e1 ∧ cont_pole ff p e2"
        by auto
      have "∃e3>0. cball p e3 ⊆ ball p e1 ∧ cont_zero ff p e3"
        when "p∈zeros"
        proof -
          define zo where "zo ≡ zorder f p"
          define zp where "zp ≡ zer_poly f p"
          define f' where "f' ≡ λw. zp w * (w - p) ^ zo"
          define ff' where "ff' ≡ (λx. deriv f' x * h x / f' x)"
          have "f holomorphic_on ball p e1"
            proof -
              have "ball p e1 ⊆ s - poles"
                using avoid_def ball_subset_cball e1_avoid that zeros_def by fastforce
              thus ?thesis using f_holo by auto
            qed
          moreover have "f p = 0" using ‹p∈zeros›
            using DiffD1 mem_Collect_eq zeros_def by blast
          moreover have "∃w∈ball p e1. f w ≠ 0"
            proof -
              define p' where "p' ≡ p+e1/2"
              have "p' ∈ ball p e1" and "p'≠p" using ‹e1>0› unfolding p'_def by (auto simp add:dist_norm)
              then show "∃w∈ball p e1. f w ≠ 0" using e1_avoid unfolding avoid_def
                apply (rule_tac x=p' in bexI)
                by (auto simp add:zeros_def)
            qed
          ultimately obtain r where
              "0 < zo" "r>0"
              "cball p r ⊆ ball p e1" and
              pp_holo:"zp holomorphic_on cball p r" and
              pp_po:"(∀w∈cball p r. f w = zp w * (w - p) ^ zo ∧ zp w ≠ 0)"
            using zorder_exist[of "ball p e1" p f,simplified,OF ‹e1>0›] unfolding zo_def zp_def
            by auto
          define e2 where "e2 ≡ r/2"
          have "e2>0" using ‹r>0› unfolding e2_def by auto
          define anal where "anal ≡ λw. deriv zp w * h w / zp w"
          define prin where "prin ≡ λw. of_nat zo * h w / (w - p)"
          have "((λw.  prin w + anal w) has_contour_integral c * zo * h p) (circlepath p e2)"
            proof (rule  has_contour_integral_add[of _ _ _ _ 0,simplified])
              have "ball p r ⊆ s"
                using ‹cball p r ⊆ ball p e1› avoid_def ball_subset_cball e1_avoid by blast
              then have "cball p e2 ⊆ s"
                using ‹r>0› unfolding e2_def by auto
              then have "(λw. of_nat zo * h w) holomorphic_on cball p e2"
                using h_holo
                by (auto intro!: holomorphic_intros)
              then show "(prin has_contour_integral c * of_nat zo * h p ) (circlepath p e2)"
                using Cauchy_integral_circlepath_simple[folded c_def, of "λw. of_nat zo * h w"]
                  ‹e2>0›
                unfolding prin_def
                by (auto simp add: mult.assoc)
              have "anal holomorphic_on ball p r" unfolding anal_def
                using pp_holo h_holo pp_po ‹ball p r ⊆ s›
                by (auto intro!: holomorphic_intros)
              then show "(anal has_contour_integral 0) (circlepath p e2)"
                using e2_def ‹r>0›
                by (auto elim!: Cauchy_theorem_disc_simple)
            qed
          then have "cont_zero ff' p e2" unfolding cont_zero_def zo_def
            proof (elim has_contour_integral_eq)
              fix w assume "w ∈ path_image (circlepath p e2)"
              then have "w∈ball p r" and "w≠p" unfolding e2_def using ‹r>0› by auto
              define wp where "wp ≡ w-p"
              have "wp≠0" and "zp w ≠0"
                unfolding wp_def using ‹w≠p› ‹w∈ball p r› pp_po by auto
              moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo"
                proof (rule DERIV_imp_deriv)
                  define der where "der ≡ zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo"
                  have po:"zo = Suc (zo - Suc 0) " using ‹zo>0› by auto
                  have "(zp has_field_derivative (deriv zp w)) (at w)"
                    using DERIV_deriv_iff_has_field_derivative pp_holo
                    by (meson Topology_Euclidean_Space.open_ball ‹w ∈ ball p r› ball_subset_cball holomorphic_derivI holomorphic_on_subset)
                  then show "(f' has_field_derivative  der) (at w)"
                    using ‹w≠p› ‹zo>0› unfolding der_def f'_def
                    by (auto intro!: derivative_eq_intros simp add:field_simps)
                qed
              ultimately show "prin w + anal w = ff' w"
                unfolding ff'_def prin_def anal_def
                apply simp
                apply (unfold f'_def)
                apply (fold wp_def)
                apply (auto simp add:field_simps)
                by (metis Suc_diff_Suc minus_nat.diff_0 power_Suc)
            qed
          then have "cont_zero ff p e2" unfolding cont_zero_def
            proof (elim has_contour_integral_eq)
              fix w assume "w ∈ path_image (circlepath p e2)"
              then have "w∈ball p r" and "w≠p" unfolding e2_def using ‹r>0› by auto
              have "deriv f' w =  deriv f w"
                proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"])
                  show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo
                    by (auto intro!: holomorphic_intros)
                next
                  have "ball p e1 - {p} ⊆ s - poles"
                    using avoid_def ball_subset_cball e1_avoid by auto
                  then have "ball p r - {p} ⊆ s - poles" using ‹cball p r ⊆ ball p e1›
                    using ball_subset_cball by blast
                  then show "f holomorphic_on ball p r - {p}" using f_holo
                    by auto
                next
                  show "open (ball p r - {p})" by auto
                next
                  show "w ∈ ball p r - {p}" using ‹w∈ball p r› ‹w≠p› by auto
                next
                  fix x assume "x ∈ ball p r - {p}"
                  then show "f' x = f x"
                    using pp_po unfolding f'_def by auto
                qed
              moreover have " f' w  =  f w "
                using ‹w ∈ ball p r› ball_subset_cball subset_iff pp_po unfolding f'_def by auto
              ultimately show "ff' w = ff w"
                unfolding ff'_def ff_def by simp
            qed
          moreover have "cball p e2 ⊆ ball p e1"
            using ‹0 < r› ‹cball p r ⊆ ball p e1› e2_def by auto
          ultimately show ?thesis using ‹e2>0› by auto
        qed
      then obtain e3 where e3:"p∈zeros ⟶ e3>0 ∧ cball p e3 ⊆ ball p e1 ∧ cont_zero ff p e3"
        by auto
      define e4 where "e4 ≡ if p∈poles then e2 else if p∈zeros then e3 else e1"
      have "e4>0" using e2 e3 ‹e1>0› unfolding e4_def by auto
      moreover have "avoid p e4" using e2 e3 ‹e1>0› e1_avoid unfolding e4_def avoid_def by auto
      moreover have "p∈poles ⟶ cont_pole ff p e4" and "p∈zeros ⟶ cont_zero ff p e4"
        by (auto simp add: e2 e3 e4_def zeros_def)
      ultimately show ?thesis by auto
    qed
  then obtain get_e where get_e:"∀p∈s. get_e p>0 ∧ avoid p (get_e p)
      ∧ (p∈poles ⟶ cont_pole ff p (get_e p)) ∧ (p∈zeros ⟶ cont_zero ff p (get_e p))"
    by metis
  define cont where "cont ≡ λp. contour_integral (circlepath p (get_e p)) ff"
  define w where "w ≡ λp. winding_number g p"
  have "contour_integral g ff = (∑p∈zeros ∪ poles. w p * cont p)"
    unfolding cont_def w_def
    proof (rule Cauchy_theorem_singularities[OF ‹open s› ‹connected s› finite _ ‹valid_path g› loop
        path_img homo])
      have "open (s - (zeros ∪ poles))" using open_Diff[OF _ finite_imp_closed[OF finite]] ‹open s› by auto
      then show "ff holomorphic_on s - (zeros ∪ poles)" unfolding ff_def using f_holo h_holo
        by (auto intro!: holomorphic_intros simp add:zeros_def)
    next
      show "∀p∈s. 0 < get_e p ∧ (∀w∈cball p (get_e p). w ∈ s ∧ (w ≠ p ⟶ w ∉ zeros ∪ poles))"
        using get_e using avoid_def by blast
    qed
  also have "... = (∑p∈zeros. w p * cont p) + (∑p∈poles. w p * cont p)"
    using finite
    apply (subst sum.union_disjoint)
    by (auto simp add:zeros_def)
  also have "... = c * ((∑p∈zeros. w p *  h p * zorder f p) - (∑p∈poles. w p *  h p * porder f p))"
    proof -
      have "(∑p∈zeros. w p * cont p) = (∑p∈zeros. c * w p *  h p * zorder f p)"
        proof (rule sum.cong[of zeros zeros,simplified])
          fix p assume "p ∈ zeros"
          show "w p * cont p = c * w p * h p * (zorder f p)"
            proof (cases "p∈s")
              assume "p ∈ s"
              have "cont p = c * h p * (zorder f p)" unfolding cont_def
                apply (rule contour_integral_unique)
                using get_e ‹p∈s› ‹p∈zeros› unfolding cont_zero_def
                by (metis mult.assoc mult.commute)
              thus ?thesis by auto
            next
              assume "p∉s"
              then have "w p=0" using homo unfolding w_def by auto
              then show ?thesis by auto
            qed
        qed
      then have "(∑p∈zeros. w p * cont p) = c * (∑p∈zeros.  w p *  h p * zorder f p)"
        apply (subst sum_distrib_left)
        by (simp add:algebra_simps)
      moreover have "(∑p∈poles. w p * cont p) = (∑p∈poles.  - c * w p *  h p * porder f p)"
        proof (rule sum.cong[of poles poles,simplified])
          fix p assume "p ∈ poles"
          show "w p * cont p = - c * w p * h p * (porder f p)"
            proof (cases "p∈s")
              assume "p ∈ s"
              have "cont p = - c * h p * (porder f p)" unfolding cont_def
                apply (rule contour_integral_unique)
                using get_e ‹p∈s› ‹p∈poles› unfolding cont_pole_def
                by (metis mult.assoc mult.commute)
              thus ?thesis by auto
            next
              assume "p∉s"
              then have "w p=0" using homo unfolding w_def by auto
              then show ?thesis by auto
            qed
        qed
      then have "(∑p∈poles. w p * cont p) = - c * (∑p∈poles. w p *  h p * porder f p)"
        apply (subst sum_distrib_left)
        by (simp add:algebra_simps)
      ultimately show ?thesis by (simp add: right_diff_distrib)
    qed
  finally show ?thesis unfolding w_def ff_def c_def by auto
qed

subsection ‹Rouche's theorem ›

theorem Rouche_theorem:
  fixes f g::"complex ⇒ complex" and s:: "complex set"
  defines "fg≡(λp. f p+ g p)"
  defines "zeros_fg≡{p. fg p =0}" and "zeros_f≡{p. f p=0}"
  assumes
    "open s" and "connected s" and
    "finite zeros_fg" and
    "finite zeros_f" and
    f_holo:"f holomorphic_on s" and
    g_holo:"g holomorphic_on s" and
    "valid_path γ" and
    loop:"pathfinish γ = pathstart γ" and
    path_img:"path_image γ ⊆ s " and
    path_less:"∀z∈path_image γ. cmod(f z) > cmod(g z)" and
    homo:"∀z. (z ∉ s) ⟶ winding_number γ z = 0"
  shows "(∑p∈zeros_fg. winding_number γ p * zorder fg p)
          = (∑p∈zeros_f. winding_number γ p * zorder f p)"
proof -
  have path_fg:"path_image γ ⊆ s - zeros_fg"
    proof -
      have False when "z∈path_image γ" and "f z + g z=0" for z
        proof -
          have "cmod (f z) > cmod (g z)" using ‹z∈path_image γ› path_less by auto
          moreover have "f z = - g z"  using ‹f z + g z =0› by (simp add: eq_neg_iff_add_eq_0)
          then have "cmod (f z) = cmod (g z)" by auto
          ultimately show False by auto
        qed
      then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto
    qed
  have path_f:"path_image γ ⊆ s - zeros_f"
    proof -
      have False when "z∈path_image γ" and "f z =0" for z
        proof -
          have "cmod (g z) < cmod (f z) " using ‹z∈path_image γ› path_less by auto
          then have "cmod (g z) < 0" using ‹f z=0› by auto
          then show False by auto
        qed
      then show ?thesis unfolding zeros_f_def using path_img by auto
    qed
  define w where "w ≡ λp. winding_number γ p"
  define c where "c ≡ 2 * complex_of_real pi * 𝗂"
  define h where "h ≡ λp. g p / f p + 1"
  obtain spikes
    where "finite spikes" and spikes: "∀x∈{0..1} - spikes. γ differentiable at x"
    using ‹valid_path γ›
    by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
  have h_contour:"((λx. deriv h x / h x) has_contour_integral 0) γ"
    proof -
      have outside_img:"0 ∈ outside (path_image (h o γ))"
        proof -
          have "h p ∈ ball 1 1" when "p∈path_image γ" for p
            proof -
              have "cmod (g p/f p) <1" using path_less[rule_format,OF that]
                apply (cases "cmod (f p) = 0")
                by (auto simp add: norm_divide)
              then show ?thesis unfolding h_def by (auto simp add:dist_complex_def)
            qed
          then have "path_image (h o γ) ⊆ ball 1 1"
            by (simp add: image_subset_iff path_image_compose)
          moreover have " (0::complex) ∉ ball 1 1" by (simp add: dist_norm)
          ultimately show "?thesis"
            using  convex_in_outside[of "ball 1 1" 0] outside_mono by blast
        qed
      have valid_h:"valid_path (h ∘ γ)"
        proof (rule valid_path_compose_holomorphic[OF ‹valid_path γ› _ _ path_f])
          show "h holomorphic_on s - zeros_f"
            unfolding h_def using f_holo g_holo
            by (auto intro!: holomorphic_intros simp add:zeros_f_def)
        next
          show "open (s - zeros_f)" using ‹finite zeros_f› ‹open s› finite_imp_closed
            by auto
        qed
      have "((λz. 1/z) has_contour_integral 0) (h ∘ γ)"
        proof -
          have "0 ∉ path_image (h ∘ γ)" using outside_img by (simp add: outside_def)
          then have "((λz. 1/z) has_contour_integral c * winding_number (h ∘ γ) 0) (h ∘ γ)"
            using has_contour_integral_winding_number[of "h o γ" 0,simplified] valid_h
            unfolding c_def by auto
          moreover have "winding_number (h o γ) 0 = 0"
            proof -
              have "0 ∈ outside (path_image (h ∘ γ))" using outside_img .
              moreover have "path (h o γ)"
                using valid_h  by (simp add: valid_path_imp_path)
              moreover have "pathfinish (h o γ) = pathstart (h o γ)"
                by (simp add: loop pathfinish_compose pathstart_compose)
              ultimately show ?thesis using winding_number_zero_in_outside by auto
            qed
          ultimately show ?thesis by auto
        qed
      moreover have "vector_derivative (h ∘ γ) (at x) = vector_derivative γ (at x) * deriv h (γ x)"
          when "x∈{0..1} - spikes" for x
        proof (rule vector_derivative_chain_at_general)
          show "γ differentiable at x" using that ‹valid_path γ› spikes by auto
        next
          define der where "der ≡ λp. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
          define t where "t ≡ γ x"
          have "f t≠0" unfolding zeros_f_def t_def
            by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that)
          moreover have "t∈s"
            using contra_subsetD path_image_def path_fg t_def that by fastforce
          ultimately have "(h has_field_derivative der t) (at t)"
            unfolding h_def der_def using g_holo f_holo ‹open s›
            by (auto intro!: holomorphic_derivI derivative_eq_intros)
          then show "h field_differentiable at (γ x)" 
            unfolding t_def field_differentiable_def by blast
        qed
      then have " (op / 1 has_contour_integral 0) (h ∘ γ)
          = ((λx. deriv h x / h x) has_contour_integral 0) γ"
        unfolding has_contour_integral
        apply (intro has_integral_spike_eq[OF negligible_finite, OF ‹finite spikes›])
        by auto
      ultimately show ?thesis by auto
    qed
  then have "contour_integral γ (λx. deriv h x / h x) = 0"
    using  contour_integral_unique by simp
  moreover have "contour_integral γ (λx. deriv fg x / fg x) = contour_integral γ (λx. deriv f x / f x)
      + contour_integral γ (λp. deriv h p / h p)"
    proof -
      have "(λp. deriv f p / f p) contour_integrable_on γ"
        proof (rule contour_integrable_holomorphic_simple[OF _ _ ‹valid_path γ› path_f])
          show "open (s - zeros_f)" using finite_imp_closed[OF ‹finite zeros_f›] ‹open s›
            by auto
          then show "(λp. deriv f p / f p) holomorphic_on s - zeros_f"
            using f_holo
            by (auto intro!: holomorphic_intros simp add:zeros_f_def)
        qed
      moreover have "(λp. deriv h p / h p) contour_integrable_on γ"
        using h_contour
        by (simp add: has_contour_integral_integrable)
      ultimately have "contour_integral γ (λx. deriv f x / f x + deriv h x / h x) =
          contour_integral γ (λp. deriv f p / f p) + contour_integral γ (λp. deriv h p / h p)"
        using contour_integral_add[of "(λp. deriv f p / f p)" γ "(λp. deriv h p / h p)" ]
        by auto
      moreover have "deriv fg p / fg p =  deriv f p / f p + deriv h p / h p"
          when "p∈ path_image γ" for p
        proof -
          have "fg p≠0" and "f p≠0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def
            by auto
          have "h p≠0"
            proof (rule ccontr)
              assume "¬ h p ≠ 0"
              then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2)
              then have "cmod (g p/f p) = 1" by auto
              moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that]
                apply (cases "cmod (f p) = 0")
                by (auto simp add: norm_divide)
              ultimately show False by auto
            qed
          have der_fg:"deriv fg p =  deriv f p + deriv g p" unfolding fg_def
            using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _  ‹open s›] path_img that
            by auto
          have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)"
            proof -
              define der where "der ≡ λp. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
              have "p∈s" using path_img that by auto
              then have "(h has_field_derivative der p) (at p)"
                unfolding h_def der_def using g_holo f_holo ‹open s› ‹f p≠0›
                by (auto intro!: derivative_eq_intros holomorphic_derivI)
              then show ?thesis unfolding der_def using DERIV_imp_deriv by auto
            qed
          show ?thesis
            apply (simp only:der_fg der_h)
            apply (auto simp add:field_simps ‹h p≠0› ‹f p≠0› ‹fg p≠0›)
            by (auto simp add:field_simps h_def ‹f p≠0› fg_def)
        qed
      then have "contour_integral γ (λp. deriv fg p / fg p)
          = contour_integral γ (λp. deriv f p / f p + deriv h p / h p)"
        by (elim contour_integral_eq)
      ultimately show ?thesis by auto
    qed
  moreover have "contour_integral γ (λx. deriv fg x / fg x) = c * (∑p∈zeros_fg. w p * zorder fg p)"
    unfolding c_def zeros_fg_def w_def
    proof (rule argument_principle[OF ‹open s› ‹connected s› _ _ ‹valid_path γ› loop _ homo
        , of _ "{}" "λ_. 1",simplified])
      show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto
      show "path_image γ ⊆ s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def .
      show " finite {p. fg p = 0}" using ‹finite zeros_fg› unfolding zeros_fg_def .
    qed
  moreover have "contour_integral γ (λx. deriv f x / f x) = c * (∑p∈zeros_f. w p * zorder f p)"
    unfolding c_def zeros_f_def w_def
    proof (rule argument_principle[OF ‹open s› ‹connected s› _ _ ‹valid_path γ› loop _ homo
        , of _ "{}" "λ_. 1",simplified])
      show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto
      show "path_image γ ⊆ s - {p. f p = 0}" using path_f unfolding zeros_f_def .
      show " finite {p. f p = 0}" using ‹finite zeros_f› unfolding zeros_f_def .
    qed
  ultimately have " c* (∑p∈zeros_fg. w p * (zorder fg p)) = c* (∑p∈zeros_f. w p * (zorder f p))"
    by auto
  then show ?thesis unfolding c_def using w_def by auto
qed


subsection ‹More facts about poles and residues›

lemma zorder_cong:
  assumes "eventually (λz. f z = g z) (nhds z)" "z = z'"
  shows   "zorder f z = zorder g z'"
proof -
  let ?P = "(λf n h r. 0 < r ∧ h holomorphic_on cball z r ∧
              (∀w∈cball z r. f w = h w * (w - z) ^ n ∧ h w ≠ 0))"
  have "(λn. n > 0 ∧ (∃h r. ?P f n h r)) = (λn. n > 0 ∧ (∃h r. ?P g n h r))"
  proof (intro ext conj_cong refl iff_exI[where ?'a = "complex ⇒ complex"], goal_cases)
    case (1 n h)
    have *: "∃r. ?P g n h r" if "∃r. ?P f n h r" and "eventually (λx. f x = g x) (nhds z)" for f g
    proof -
      from that(1) obtain r where "?P f n h r" by blast
      moreover from that(2) obtain r' where "r' > 0" "⋀w. dist w z < r' ⟹ f w = g w"
        by (auto simp: eventually_nhds_metric)
      hence "∀w∈cball z (r'/2). f w = g w" by (auto simp: dist_commute)
      ultimately show ?thesis using ‹r' > 0›
        by (intro exI[of _ "min r (r'/2)"]) (auto simp: cball_def)
    qed
    from assms have eq': "eventually (λz. g z = f z) (nhds z)"
      by (simp add: eq_commute)
    show ?case
      by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']])
  qed
  with assms(2) show ?thesis unfolding zorder_def by simp
qed

lemma porder_cong:
  assumes "eventually (λz. f z = g z) (at z)" "z = z'"
  shows   "porder f z = porder g z'"
proof -
  from assms(1) have *: "eventually (λw. w ≠ z ⟶ f w = g w) (nhds z)"
    by (auto simp: eventually_at_filter)
  from assms(2) show ?thesis
    unfolding porder_def Let_def
    by (intro zorder_cong eventually_mono [OF *]) auto
qed

lemma zer_poly_cong:
  assumes "eventually (λz. f z = g z) (nhds z)" "z = z'"
  shows   "zer_poly f z = zer_poly g z'"
  unfolding zer_poly_def
proof (rule Eps_cong, goal_cases)
  case (1 h)
  let ?P = "λw f. f w = h w * (w - z) ^ zorder f z ∧ h w ≠ 0"
  from assms have eq': "eventually (λz. g z = f z) (nhds z)"
    by (simp add: eq_commute)
  have "∃r>0. h holomorphic_on cball z r ∧ (∀w∈cball z r. ?P w g)"
    if "r > 0" "h holomorphic_on cball z r" "⋀w. w ∈ cball z r ⟹ ?P w f"
       "eventually (λz. f z = g z) (nhds z)" for f g r
  proof -
    from that have [simp]: "zorder f z = zorder g z"
      by (intro zorder_cong) auto
    from that(4) obtain r' where r': "r' > 0" "⋀w. w ∈ ball z r' ⟹ g w = f w"
      by (auto simp: eventually_nhds_metric ball_def dist_commute)
    define R where "R = min r (r' / 2)"
    have "R > 0" "cball z R ⊆ cball z r" "cball z R ⊆ ball z r'"
      using that(1) r' by (auto simp: R_def)
    with that(1,2,3) r'
    have "R > 0" "h holomorphic_on cball z R" "∀w∈cball z R. ?P w g" 
      by force+
    thus ?thesis by blast
  qed
  from this[of _ f g] and this[of _ g f] and assms and eq'
    show ?case by blast
qed
      
lemma pol_poly_cong:
  assumes "eventually (λz. f z = g z) (at z)" "z = z'"
  shows   "pol_poly f z = pol_poly g z'"
proof -
  from assms have *: "eventually (λw. w ≠ z ⟶ f w = g w) (nhds z)"
    by (auto simp: eventually_at_filter)
  have "zer_poly (λx. if x = z then 0 else inverse (f x)) z =
          zer_poly (λx. if x = z' then 0 else inverse (g x)) z"
    by (intro zer_poly_cong eventually_mono[OF *] refl) (auto simp: assms(2))
  thus "pol_poly f z = pol_poly g z'"
    by (simp add: pol_poly_def Let_def o_def fun_eq_iff assms(2))
qed

lemma porder_nonzero_div_power:
  assumes "open s" "z ∈ s" "f holomorphic_on s" "f z ≠ 0" "n > 0"
  shows   "porder (λw. f w / (w - z) ^ n) z = n"
proof -
  let ?s' = "(f -` (-{0}) ∩ s)"
  have "continuous_on s f"
    by (rule holomorphic_on_imp_continuous_on) fact
  moreover have "open (-{0::complex})" by auto
  ultimately have s': "open ?s'"
    unfolding continuous_on_open_vimage[OF ‹open s›] by blast
  show ?thesis unfolding Let_def porder_def
  proof (rule zorder_eqI)
    show "(λx. inverse (f x)) holomorphic_on ?s'"
      using assms by (auto intro!: holomorphic_intros)
  qed (insert assms s', auto simp: field_simps)
qed

lemma is_pole_inverse_power: "n > 0 ⟹ is_pole (λz::complex. 1 / (z - a) ^ n) a"
  unfolding is_pole_def inverse_eq_divide [symmetric]
  by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros)
     (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros)

lemma is_pole_inverse: "is_pole (λz::complex. 1 / (z - a)) a"
  using is_pole_inverse_power[of 1 a] by simp

lemma is_pole_divide:
  fixes f :: "'a :: t2_space ⇒ 'b :: real_normed_field"
  assumes "isCont f z" "filterlim g (at 0) (at z)" "f z ≠ 0"
  shows   "is_pole (λz. f z / g z) z"
proof -
  have "filterlim (λz. f z * inverse (g z)) at_infinity (at z)"
    by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"]
                 filterlim_compose[OF filterlim_inverse_at_infinity])+
       (insert assms, auto simp: isCont_def)
  thus ?thesis by (simp add: divide_simps is_pole_def)
qed

lemma is_pole_basic:
  assumes "f holomorphic_on A" "open A" "z ∈ A" "f z ≠ 0" "n > 0"
  shows   "is_pole (λw. f w / (w - z) ^ n) z"
proof (rule is_pole_divide)
  have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact
  with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at)
  have "filterlim (λw. (w - z) ^ n) (nhds 0) (at z)"
    using assms by (auto intro!: tendsto_eq_intros)
  thus "filterlim (λw. (w - z) ^ n) (at 0) (at z)"
    by (intro filterlim_atI tendsto_eq_intros)
       (insert assms, auto simp: eventually_at_filter)
qed fact+

lemma is_pole_basic':
  assumes "f holomorphic_on A" "open A" "0 ∈ A" "f 0 ≠ 0" "n > 0"
  shows   "is_pole (λw. f w / w ^ n) 0"
  using is_pole_basic[of f A 0] assms by simp

lemma zer_poly_eq:
  assumes "open s" "connected s" "z ∈ s" "f holomorphic_on s" "f z = 0" "∃w∈s. f w ≠ 0"
  shows "eventually (λw. zer_poly f z w = f w / (w - z) ^ zorder f z) (at z)"
proof -
  from zorder_exist [OF assms] obtain r where r: "r > 0" 
    and "∀w∈cball z r. f w = zer_poly f z w * (w - z) ^ zorder f z" by blast
  hence *: "∀w∈ball z r - {z}. zer_poly f z w = f w / (w - z) ^ zorder f z" 
    by (auto simp: field_simps)
  have "eventually (λw. w ∈ ball z r - {z}) (at z)"
    using r eventually_at_ball'[of r z UNIV] by auto
  thus ?thesis by eventually_elim (insert *, auto)
qed

lemma pol_poly_eq:
  assumes "open s" "z ∈ s" "f holomorphic_on s - {z}" "is_pole f z" "∃w∈s. f w ≠ 0"
  shows "eventually (λw. pol_poly f z w = f w * (w - z) ^ porder f z) (at z)"
proof -
  from porder_exist[OF assms(1-4)] obtain r where r: "r > 0" 
    and "∀w∈cball z r. w ≠ z ⟶ f w = pol_poly f z w / (w - z) ^ porder f z" by blast
  hence *: "∀w∈ball z r - {z}. pol_poly f z w = f w * (w - z) ^ porder f z" 
    by (auto simp: field_simps)
  have "eventually (λw. w ∈ ball z r - {z}) (at z)"
    using r eventually_at_ball'[of r z UNIV] by auto
  thus ?thesis by eventually_elim (insert *, auto)
qed

lemma lhopital_complex_simple:
  assumes "(f has_field_derivative f') (at z)" 
  assumes "(g has_field_derivative g') (at z)"
  assumes "f z = 0" "g z = 0" "g' ≠ 0" "f' / g' = c"
  shows   "((λw. f w / g w) ⤏ c) (at z)"
proof -
  have "eventually (λw. w ≠ z) (at z)"
    by (auto simp: eventually_at_filter)
  hence "eventually (λw. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)"
    by eventually_elim (simp add: assms divide_simps)
  moreover have "((λw. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) ⤏ f' / g') (at z)"
    by (intro tendsto_divide has_field_derivativeD assms)
  ultimately have "((λw. f w / g w) ⤏ f' / g') (at z)"
    by (rule Lim_transform_eventually)
  with assms show ?thesis by simp
qed

lemma porder_eqI:
  assumes "open s" "z ∈ s" "g holomorphic_on s" "g z ≠ 0" "n > 0"
  assumes "⋀w. w ∈ s - {z} ⟹ f w = g w / (w - z) ^ n"
  shows   "porder f z = n"
proof -
  define f' where "f' = (λx. if x = z then 0 else inverse (f x))"
  define g' where "g' = (λx. inverse (g x))"
  define s' where "s' = (g -` (-{0}) ∩ s)"
  have "continuous_on s g"
    by (intro holomorphic_on_imp_continuous_on) fact
  hence "open s'"
    unfolding s'_def using assms by (subst (asm) continuous_on_open_vimage) blast+
  have s': "z ∈ s'" "g' holomorphic_on s'" "g' z ≠ 0" using assms 
    by (auto simp: s'_def g'_def intro!: holomorphic_intros)
  have f'_g': "f' w = g' w * (w - z) ^ n" if "w ∈ s'" for w
    unfolding f'_def g'_def using that ‹n > 0›
    by (auto simp: assms(6) field_simps s'_def)
  have "porder f z = zorder f' z"
    by (simp add: porder_def f'_def)
  also have "… = n" using assms f'_g' 
    by (intro zorder_eqI[OF ‹open s'› s']) (auto simp: f'_def g'_def field_simps s'_def)
  finally show ?thesis .
qed

lemma simple_poleI':
  assumes "open s" "connected s" "z ∈ s"
  assumes "⋀w. w ∈ s - {z} ⟹ 
             ((λw. inverse (f w)) has_field_derivative f' w) (at w)"
  assumes "f holomorphic_on s - {z}" "f' holomorphic_on s" "is_pole f z" "f' z ≠ 0"
  shows   "porder f z = 1"
proof -
  define g where "g = (λw. if w = z then 0 else inverse (f w))"
  from ‹is_pole f z› have "eventually (λw. f w ≠ 0) (at z)"
    unfolding is_pole_def using filterlim_at_infinity_imp_eventually_ne by blast
  then obtain s'' where s'': "open s''" "z ∈ s''" "∀w∈s''-{z}. f w ≠ 0"
    by (auto simp: eventually_at_topological)
  from assms(1) and s''(1) have "open (s ∩ s'')" by auto
  then obtain r where r: "r > 0" "ball z r ⊆ s ∩ s''"
    using assms(3) s''(2) by (subst (asm) open_contains_ball) blast
  define s' where "s' = ball z r"
  hence s': "open s'" "connected s'" "z ∈ s'" "s' ⊆ s" "∀w∈s'-{z}. f w ≠ 0"
    using r s'' by (auto simp: s'_def)
  have s'_ne: "s' - {z} ≠ {}"
    using r unfolding s'_def by (intro ball_minus_countable_nonempty) auto

  have "porder f z = zorder g z"
    by (simp add: porder_def g_def)
  also have "… = 1"
  proof (rule simple_zeroI')
    fix w assume w: "w ∈ s'"
    have [holomorphic_intros]: "g holomorphic_on s'" unfolding g_def using assms s'
      by (intro is_pole_inverse_holomorphic holomorphic_on_subset[OF assms(5)]) auto
    hence "(g has_field_derivative deriv g w) (at w)"
      using w s' by (intro holomorphic_derivI)
    also have deriv_g: "deriv g w = f' w" if "w ∈ s' - {z}" for w
    proof -
      from that have ne: "eventually (λw. w ≠ z) (nhds w)"
        by (intro t1_space_nhds) auto
      have "deriv g w = deriv (λw. inverse (f w)) w"
        by (intro deriv_cong_ev refl eventually_mono [OF ne]) (auto simp: g_def)
      also from assms(4)[of w] that s' have "… = f' w"
        by (auto dest: DERIV_imp_deriv)
      finally show ?thesis .
    qed
    have "deriv g w = f' w"
      by (rule analytic_continuation_open[of "s' - {z}" s' "deriv g" f'])
         (insert s' assms s'_ne deriv_g w, 
          auto intro!: holomorphic_intros holomorphic_on_subset[OF assms(6)])
    finally show "(g has_field_derivative f' w) (at w)" .
  qed (insert assms s', auto simp: g_def)
  finally show ?thesis .
qed

lemma residue_holomorphic_over_power:
  assumes "open A" "z0 ∈ A" "f holomorphic_on A"
  shows   "residue (λz. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n"
proof -
  let ?f = "λz. f z / (z - z0) ^ Suc n"
  from assms(1,2) obtain r where r: "r > 0" "cball z0 r ⊆ A"
    by (auto simp: open_contains_cball)
  have "(?f has_contour_integral 2 * pi * 𝗂 * residue ?f z0) (circlepath z0 r)"
    using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros)
  moreover have "(?f has_contour_integral 2 * pi * 𝗂 / fact n * (deriv ^^ n) f z0) (circlepath z0 r)"
    using assms r
    by (intro Cauchy_has_contour_integral_higher_derivative_circlepath)
       (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on)
  ultimately have "2 * pi * 𝗂 * residue ?f z0 = 2 * pi * 𝗂 / fact n * (deriv ^^ n) f z0"  
    by (rule has_contour_integral_unique)
  thus ?thesis by (simp add: field_simps)
qed

lemma residue_holomorphic_over_power':
  assumes "open A" "0 ∈ A" "f holomorphic_on A"
  shows   "residue (λz. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n"
  using residue_holomorphic_over_power[OF assms] by simp

lemma zer_poly_eqI:
  fixes f :: "complex ⇒ complex" and z0 :: complex
  defines "n ≡ zorder f z0"
  assumes "open A" "connected A" "z0 ∈ A" "f holomorphic_on A" "f z0 = 0" "∃z∈A. f z ≠ 0"
  assumes lim: "((λx. f (g x) / (g x - z0) ^ n) ⤏ c) F"
  assumes g: "filterlim g (at z0) F" and "F ≠ bot"
  shows   "zer_poly f z0 z0 = c"
proof -
  from zorder_exist[OF assms(2-7)] obtain r where
    r: "r > 0" "cball z0 r ⊆ A" "zer_poly f z0 holomorphic_on cball z0 r"
       "⋀w. w ∈ cball z0 r ⟹ f w = zer_poly f z0 w * (w - z0) ^ n"
    unfolding n_def by blast
  from r(1) have "eventually (λw. w ∈ ball z0 r ∧ w ≠ z0) (at z0)"
    using eventually_at_ball'[of r z0 UNIV] by auto
  hence "eventually (λw. zer_poly f z0 w = f w / (w - z0) ^ n) (at z0)"
    by eventually_elim (insert r, auto simp: field_simps)
  moreover have "continuous_on (ball z0 r) (zer_poly f z0)"
    using r by (intro holomorphic_on_imp_continuous_on) auto
  with r(1,2) have "isCont (zer_poly f z0) z0"
    by (auto simp: continuous_on_eq_continuous_at)
  hence "(zer_poly f z0 ⤏ zer_poly f z0 z0) (at z0)"
    unfolding isCont_def .
  ultimately have "((λw. f w / (w - z0) ^ n) ⤏ zer_poly f z0 z0) (at z0)"
    by (rule Lim_transform_eventually)
  hence "((λx. f (g x) / (g x - z0) ^ n) ⤏ zer_poly f z0 z0) F"
    by (rule filterlim_compose[OF _ g])
  from tendsto_unique[OF ‹F ≠ bot› this lim] show ?thesis .
qed

lemma pol_poly_eqI:
  fixes f :: "complex ⇒ complex" and z0 :: complex
  defines "n ≡ porder f z0"
  assumes "open A" "z0 ∈ A" "f holomorphic_on A-{z0}" "is_pole f z0"
  assumes lim: "((λx. f (g x) * (g x - z0) ^ n) ⤏ c) F"
  assumes g: "filterlim g (at z0) F" and "F ≠ bot"
  shows   "pol_poly f z0 z0 = c"
proof -
  from porder_exist[OF assms(2-5)] obtain r where
    r: "r > 0" "cball z0 r ⊆ A" "pol_poly f z0 holomorphic_on cball z0 r"
       "⋀w. w ∈ cball z0 r - {z0} ⟹ f w = pol_poly f z0 w / (w - z0) ^ n"
    unfolding n_def by blast
  from r(1) have "eventually (λw. w ∈ ball z0 r ∧ w ≠ z0) (at z0)"
    using eventually_at_ball'[of r z0 UNIV] by auto
  hence "eventually (λw. pol_poly f z0 w = f w * (w - z0) ^ n) (at z0)"
    by eventually_elim (insert r, auto simp: field_simps)
  moreover have "continuous_on (ball z0 r) (pol_poly f z0)"
    using r by (intro holomorphic_on_imp_continuous_on) auto
  with r(1,2) have "isCont (pol_poly f z0) z0"
    by (auto simp: continuous_on_eq_continuous_at)
  hence "(pol_poly f z0 ⤏ pol_poly f z0 z0) (at z0)"
    unfolding isCont_def .
  ultimately have "((λw. f w * (w - z0) ^ n) ⤏ pol_poly f z0 z0) (at z0)"
    by (rule Lim_transform_eventually)
  hence "((λx. f (g x) * (g x - z0) ^ n) ⤏ pol_poly f z0 z0) F"
    by (rule filterlim_compose[OF _ g])
  from tendsto_unique[OF ‹F ≠ bot› this lim] show ?thesis .
qed

lemma residue_simple_pole:
  assumes "open A" "z0 ∈ A" "f holomorphic_on A - {z0}" 
  assumes "is_pole f z0" "porder f z0 = 1"
  shows   "residue f z0 = pol_poly f z0 z0"
  using assms by (subst residue_porder[of A]) simp_all

lemma residue_simple_pole_limit:
  assumes "open A" "z0 ∈ A" "f holomorphic_on A - {z0}" 
  assumes "is_pole f z0" "porder f z0 = 1"
  assumes "((λx. f (g x) * (g x - z0)) ⤏ c) F"
  assumes "filterlim g (at z0) F" "F ≠ bot"
  shows   "residue f z0 = c"
proof -
  have "residue f z0 = pol_poly f z0 z0"
    by (rule residue_simple_pole assms)+
  also have "… = c"
    using assms by (intro pol_poly_eqI[of A z0 f g c F]) auto
  finally show ?thesis .
qed

(* TODO: This is a mess and could be done much more easily if we had
   a nice compositional theory of poles and zeros *)
lemma
  assumes "open s" "connected s" "z ∈ s" "f holomorphic_on s" "g holomorphic_on s"
  assumes "(g has_field_derivative g') (at z)"
  assumes "f z ≠ 0" "g z = 0" "g' ≠ 0"
  shows   porder_simple_pole_deriv: "porder (λw. f w / g w) z = 1"
    and   residue_simple_pole_deriv: "residue (λw. f w / g w) z = f z / g'"
proof -
  have "∃w∈s. g w ≠ 0"
  proof (rule ccontr)
    assume *: "¬(∃w∈s. g w ≠ 0)"
    have **: "eventually (λw. w ∈ s) (nhds z)"
      by (intro eventually_nhds_in_open assms)
    from * have "deriv g z = deriv (λ_. 0) z"
      by (intro deriv_cong_ev eventually_mono [OF **]) auto
    also have "… = 0" by simp
    also from assms have "deriv g z = g'" by (auto dest: DERIV_imp_deriv)
    finally show False using ‹g' ≠ 0› by contradiction
  qed
  then obtain w where w: "w ∈ s" "g w ≠ 0" by blast
  from isolated_zeros[OF assms(5) assms(1-3,8) w]
  obtain r where r: "r > 0" "ball z r ⊆ s" "⋀w. w ∈ ball z r - {z} ⟹ g w ≠ 0"
    by blast
  from assms r have holo: "(λw. f w / g w) holomorphic_on ball z r - {z}"
    by (auto intro!: holomorphic_intros)

  have "eventually (λw. w ∈ ball z r - {z}) (at z)"
    using eventually_at_ball'[OF r(1), of z UNIV] by auto
  hence "eventually (λw. g w ≠ 0) (at z)"
    by eventually_elim (use r in auto)
  moreover have "continuous_on s g" 
    by (intro holomorphic_on_imp_continuous_on) fact
  with assms have "isCont g z"
    by (auto simp: continuous_on_eq_continuous_at)
  ultimately have "filterlim g (at 0) (at z)"
    using ‹g z = 0› by (auto simp: filterlim_at isCont_def)
  moreover have "continuous_on s f" by (intro holomorphic_on_imp_continuous_on) fact
  with assms have "isCont f z"
    by (auto simp: continuous_on_eq_continuous_at)
  ultimately have pole: "is_pole (λw. f w / g w) z" 
    unfolding is_pole_def using ‹f z ≠ 0›
    by (intro filterlim_divide_at_infinity[of _ "f z"]) (auto simp: isCont_def)

  have "continuous_on s f" by (intro holomorphic_on_imp_continuous_on) fact
  moreover have "open (-{0::complex})" by auto
  ultimately have "open (f -` (-{0}) ∩ s)" using ‹open s›
    by (subst (asm) continuous_on_open_vimage) blast+
  moreover have "z ∈ f -` (-{0}) ∩ s" using assms by auto
  ultimately obtain r' where r': "r' > 0" "ball z r' ⊆ f -` (-{0}) ∩ s"
    unfolding open_contains_ball by blast

  let ?D = "λw. (f w * deriv g w - g w * deriv f w) / f w ^ 2"
  show "porder (λw. f w / g w) z = 1"
  proof (rule simple_poleI')
    show "open (ball z (min r r'))" "connected (ball z (min r r'))" "z ∈ ball z (min r r')"
      using r'(1) r(1) by auto
  next
    fix w assume "w ∈ ball z (min r r') - {z}"
    with r' have "w ∈ s" "f w ≠ 0" by auto
    have "((λw. g w / f w) has_field_derivative ?D w) (at w)"
      by (rule derivative_eq_intros holomorphic_derivI[OF assms(4)] 
            holomorphic_derivI[OF assms(5)] | fact)+ 
         (simp_all add: algebra_simps power2_eq_square)
    thus "((λw. inverse (f w / g w)) has_field_derivative ?D w) (at w)"
      by (simp add: divide_simps)
  next
    from r' show "?D holomorphic_on ball z (min r r')"
      by (intro holomorphic_intros holomorphic_on_subset[OF assms(4)]
                holomorphic_on_subset[OF assms(5)]) auto
  next
    from assms have "deriv g z = g'"
      by (auto dest: DERIV_imp_deriv)
    with assms r' show "(f z * deriv g z - g z * deriv f z) / (f z)2 ≠ 0"
      by simp
  qed (insert pole holo, auto)

  show "residue (λw. f w / g w) z = f z / g'"
  proof (rule residue_simple_pole_limit)
    show "porder (λw. f w / g w) z = 1" by fact
    from r show "open (ball z r)" "z ∈ ball z r" by auto
    
    have "((λw. (f w * (w - z)) / g w) ⤏ f z / g') (at z)"
    proof (rule lhopital_complex_simple)
      show "((λw. f w * (w - z)) has_field_derivative f z) (at z)"
        using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF assms(4)])
      show "(g has_field_derivative g') (at z)" by fact
    qed (insert assms, auto)
    thus "((λw. (f w / g w) * (w - z)) ⤏ f z / g') (at z)"
      by (simp add: divide_simps)
  qed (insert holo pole, auto simp: filterlim_ident)
qed

end