src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy
changeset 72266 1e02b86eb517
parent 71201 6617fb368a06
child 72379 504fe7365820
equal deleted inserted replaced
72259:25cf074a4188 72266:1e02b86eb517
     4 begin
     4 begin
     5 
     5 
     6 subsection\<open>Proof\<close>
     6 subsection\<open>Proof\<close>
     7 
     7 
     8 lemma Cauchy_integral_formula_weak:
     8 lemma Cauchy_integral_formula_weak:
     9     assumes s: "convex s" and "finite k" and conf: "continuous_on s f"
     9     assumes S: "convex S" and "finite k" and conf: "continuous_on S f"
    10         and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)"
    10         and fcd: "(\<And>x. x \<in> interior S - k \<Longrightarrow> f field_differentiable at x)"
    11         and z: "z \<in> interior s - k" and vpg: "valid_path \<gamma>"
    11         and z: "z \<in> interior S - k" and vpg: "valid_path \<gamma>"
    12         and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
    12         and pasz: "path_image \<gamma> \<subseteq> S - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
    13       shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
    13       shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
    14 proof -
    14 proof -
       
    15   let ?fz = "\<lambda>w. (f w - f z)/(w - z)"
    15   obtain f' where f': "(f has_field_derivative f') (at z)"
    16   obtain f' where f': "(f has_field_derivative f') (at z)"
    16     using fcd [OF z] by (auto simp: field_differentiable_def)
    17     using fcd [OF z] by (auto simp: field_differentiable_def)
    17   have pas: "path_image \<gamma> \<subseteq> s" and znotin: "z \<notin> path_image \<gamma>" using pasz by blast+
    18   have pas: "path_image \<gamma> \<subseteq> S" and znotin: "z \<notin> path_image \<gamma>" using pasz by blast+
    18   have c: "continuous (at x within s) (\<lambda>w. if w = z then f' else (f w - f z) / (w - z))" if "x \<in> s" for x
    19   have c: "continuous (at x within S) (\<lambda>w. if w = z then f' else (f w - f z) / (w - z))" if "x \<in> S" for x
    19   proof (cases "x = z")
    20   proof (cases "x = z")
    20     case True then show ?thesis
    21     case True then show ?thesis
    21       apply (simp add: continuous_within)
    22       using LIM_equal [of "z" ?fz "\<lambda>w. if w = z then f' else ?fz w"] has_field_derivativeD [OF f'] 
    22       apply (rule Lim_transform_away_within [of _ "z+1" _ "\<lambda>w::complex. (f w - f z)/(w - z)"])
    23       by (force simp add: continuous_within Lim_at_imp_Lim_at_within)
    23       using has_field_derivative_at_within has_field_derivative_iff f'
       
    24       apply (fastforce simp add:)+
       
    25       done
       
    26   next
    24   next
    27     case False
    25     case False
    28     then have dxz: "dist x z > 0" by auto
    26     then have dxz: "dist x z > 0" by auto
    29     have cf: "continuous (at x within s) f"
    27     have cf: "continuous (at x within S) f"
    30       using conf continuous_on_eq_continuous_within that by blast
    28       using conf continuous_on_eq_continuous_within that by blast
    31     have "continuous (at x within s) (\<lambda>w. (f w - f z) / (w - z))"
    29     have "continuous (at x within S) (\<lambda>w. (f w - f z) / (w - z))"
    32       by (rule cf continuous_intros | simp add: False)+
    30       by (rule cf continuous_intros | simp add: False)+
    33     then show ?thesis
    31     then show ?thesis
    34       apply (rule continuous_transform_within [OF _ dxz that, of "\<lambda>w::complex. (f w - f z)/(w - z)"])
    32       apply (rule continuous_transform_within [OF _ dxz that, of ?fz])
    35       apply (force simp: dist_commute)
    33       apply (force simp: dist_commute)
    36       done
    34       done
    37   qed
    35   qed
    38   have fink': "finite (insert z k)" using \<open>finite k\<close> by blast
    36   have fink': "finite (insert z k)" using \<open>finite k\<close> by blast
    39   have *: "((\<lambda>w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>"
    37   have *: "((\<lambda>w. if w = z then f' else ?fz w) has_contour_integral 0) \<gamma>"
    40     apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop])
    38   proof (rule Cauchy_theorem_convex [OF _ S fink' _ vpg pas loop])
    41     using c apply (force simp: continuous_on_eq_continuous_within)
    39     show "(\<lambda>w. if w = z then f' else ?fz w) field_differentiable at w" 
    42     apply (rename_tac w)
    40       if "w \<in> interior S - insert z k" for w
    43     apply (rule_tac d="dist w z" and f = "\<lambda>w. (f w - f z)/(w - z)" in field_differentiable_transform_within)
    41     proof (rule field_differentiable_transform_within)
    44     apply (simp_all add: dist_pos_lt dist_commute)
    42       show "(\<lambda>w. ?fz w) field_differentiable at w"
    45     apply (metis less_irrefl)
    43         using that by (intro derivative_intros fcd; simp)
    46     apply (rule derivative_intros fcd | simp)+
    44     qed (use that in \<open>auto simp add: dist_pos_lt dist_commute\<close>)
    47     done
    45   qed (use c in \<open>force simp: continuous_on_eq_continuous_within\<close>)
    48   show ?thesis
    46   show ?thesis
    49     apply (rule has_contour_integral_eq)
    47     apply (rule has_contour_integral_eq)
    50     using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *]
    48     using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *]
    51     apply (auto simp: ac_simps divide_simps)
    49     apply (auto simp: ac_simps divide_simps)
    52     done
    50     done
    53 qed
    51 qed
    54 
    52 
    55 theorem Cauchy_integral_formula_convex_simple:
    53 theorem Cauchy_integral_formula_convex_simple:
    56     "\<lbrakk>convex s; f holomorphic_on s; z \<in> interior s; valid_path \<gamma>; path_image \<gamma> \<subseteq> s - {z};
    54   assumes "convex S" and holf: "f holomorphic_on S" and "z \<in> interior S" "valid_path \<gamma>" "path_image \<gamma> \<subseteq> S - {z}"
    57       pathfinish \<gamma> = pathstart \<gamma>\<rbrakk>
    55       "pathfinish \<gamma> = pathstart \<gamma>"
    58      \<Longrightarrow> ((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
    56     shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
    59   apply (rule Cauchy_integral_formula_weak [where k = "{}"])
    57 proof -
    60   using holomorphic_on_imp_continuous_on
    58   have "\<And>x. x \<in> interior S \<Longrightarrow> f field_differentiable at x"
    61   by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE)
    59     using holf at_within_interior holomorphic_onD interior_subset by fastforce
       
    60   then show ?thesis
       
    61     using assms
       
    62     by (intro Cauchy_integral_formula_weak [where k = "{}"]) (auto simp: holomorphic_on_imp_continuous_on)
       
    63 qed
    62 
    64 
    63 text\<open> Hence the Cauchy formula for points inside a circle.\<close>
    65 text\<open> Hence the Cauchy formula for points inside a circle.\<close>
    64 
    66 
    65 theorem Cauchy_integral_circlepath:
    67 theorem Cauchy_integral_circlepath:
    66   assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r"
    68   assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r"
    69 proof -
    71 proof -
    70   have "r > 0"
    72   have "r > 0"
    71     using assms le_less_trans norm_ge_zero by blast
    73     using assms le_less_trans norm_ge_zero by blast
    72   have "((\<lambda>u. f u / (u - w)) has_contour_integral (2 * pi) * \<i> * winding_number (circlepath z r) w * f w)
    74   have "((\<lambda>u. f u / (u - w)) has_contour_integral (2 * pi) * \<i> * winding_number (circlepath z r) w * f w)
    73         (circlepath z r)"
    75         (circlepath z r)"
    74   proof (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"])
    76   proof (rule Cauchy_integral_formula_weak [where S = "cball z r" and k = "{}"])
    75     show "\<And>x. x \<in> interior (cball z r) - {} \<Longrightarrow>
    77     show "\<And>x. x \<in> interior (cball z r) - {} \<Longrightarrow>
    76          f field_differentiable at x"
    78          f field_differentiable at x"
    77       using holf holomorphic_on_imp_differentiable_at by auto
    79       using holf holomorphic_on_imp_differentiable_at by auto
    78     have "w \<notin> sphere z r"
    80     have "w \<notin> sphere z r"
    79       by simp (metis dist_commute dist_norm not_le order_refl wz)
    81       by simp (metis dist_commute dist_norm not_le order_refl wz)
    93 subsection\<^marker>\<open>tag unimportant\<close> \<open>General stepping result for derivative formulas\<close>
    95 subsection\<^marker>\<open>tag unimportant\<close> \<open>General stepping result for derivative formulas\<close>
    94 
    96 
    95 lemma Cauchy_next_derivative:
    97 lemma Cauchy_next_derivative:
    96   assumes "continuous_on (path_image \<gamma>) f'"
    98   assumes "continuous_on (path_image \<gamma>) f'"
    97       and leB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
    99       and leB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
    98       and int: "\<And>w. w \<in> s - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f' u / (u - w)^k) has_contour_integral f w) \<gamma>"
   100       and int: "\<And>w. w \<in> S - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f' u / (u - w)^k) has_contour_integral f w) \<gamma>"
    99       and k: "k \<noteq> 0"
   101       and k: "k \<noteq> 0"
   100       and "open s"
   102       and "open S"
   101       and \<gamma>: "valid_path \<gamma>"
   103       and \<gamma>: "valid_path \<gamma>"
   102       and w: "w \<in> s - path_image \<gamma>"
   104       and w: "w \<in> S - path_image \<gamma>"
   103     shows "(\<lambda>u. f' u / (u - w)^(Suc k)) contour_integrable_on \<gamma>"
   105     shows "(\<lambda>u. f' u / (u - w)^(Suc k)) contour_integrable_on \<gamma>"
   104       and "(f has_field_derivative (k * contour_integral \<gamma> (\<lambda>u. f' u/(u - w)^(Suc k))))
   106       and "(f has_field_derivative (k * contour_integral \<gamma> (\<lambda>u. f' u/(u - w)^(Suc k))))
   105            (at w)"  (is "?thes2")
   107            (at w)"  (is "?thes2")
   106 proof -
   108 proof -
   107   have "open (s - path_image \<gamma>)" using \<open>open s\<close> closed_valid_path_image \<gamma> by blast
   109   have "open (S - path_image \<gamma>)" using \<open>open S\<close> closed_valid_path_image \<gamma> by blast
   108   then obtain d where "d>0" and d: "ball w d \<subseteq> s - path_image \<gamma>" using w
   110   then obtain d where "d>0" and d: "ball w d \<subseteq> S - path_image \<gamma>" using w
   109     using open_contains_ball by blast
   111     using open_contains_ball by blast
   110   have [simp]: "\<And>n. cmod (1 + of_nat n) = 1 + of_nat n"
   112   have [simp]: "\<And>n. cmod (1 + of_nat n) = 1 + of_nat n"
   111     by (metis norm_of_nat of_nat_Suc)
   113     by (metis norm_of_nat of_nat_Suc)
   112   have cint: "\<And>x. \<lbrakk>x \<noteq> w; cmod (x - w) < d\<rbrakk>
   114   have cint: "\<And>x. \<lbrakk>x \<noteq> w; cmod (x - w) < d\<rbrakk>
   113          \<Longrightarrow> (\<lambda>z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \<gamma>"
   115          \<Longrightarrow> (\<lambda>z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \<gamma>"
   114     apply (rule contour_integrable_div [OF contour_integrable_diff])
       
   115     using int w d
   116     using int w d
   116     by (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+
   117     apply (intro contour_integrable_div contour_integrable_diff has_contour_integral_integrable)
       
   118     by (force simp: dist_norm norm_minus_commute)
   117   have 1: "\<forall>\<^sub>F n in at w. (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k)
   119   have 1: "\<forall>\<^sub>F n in at w. (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k)
   118                          contour_integrable_on \<gamma>"
   120                          contour_integrable_on \<gamma>"
   119     unfolding eventually_at
   121     unfolding eventually_at
   120     apply (rule_tac x=d in exI)
   122     apply (rule_tac x=d in exI)
   121     apply (simp add: \<open>d > 0\<close> dist_norm field_simps cint)
   123     apply (simp add: \<open>d > 0\<close> dist_norm field_simps cint)
   258     by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto
   260     by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto
   259   have **: "contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) =
   261   have **: "contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) =
   260               (f u - f w) / (u - w) / k"
   262               (f u - f w) / (u - w) / k"
   261     if "dist u w < d" for u
   263     if "dist u w < d" for u
   262   proof -
   264   proof -
   263     have u: "u \<in> s - path_image \<gamma>"
   265     have u: "u \<in> S - path_image \<gamma>"
   264       by (metis subsetD d dist_commute mem_ball that)
   266       by (metis subsetD d dist_commute mem_ball that)
       
   267     have \<section>: "((\<lambda>x. f' x * inverse (x - u) ^ k) has_contour_integral f u) \<gamma>"
       
   268             "((\<lambda>x. f' x * inverse (x - w) ^ k) has_contour_integral f w) \<gamma>"
       
   269       using u w by (simp_all add: field_simps int)
   265     show ?thesis
   270     show ?thesis
   266       apply (rule contour_integral_unique)
   271       apply (rule contour_integral_unique)
   267       apply (simp add: diff_divide_distrib algebra_simps)
   272       apply (simp add: diff_divide_distrib algebra_simps \<section> has_contour_integral_diff has_contour_integral_div)
   268       apply (intro has_contour_integral_diff has_contour_integral_div)
       
   269       using u w apply (simp_all add: field_simps int)
       
   270       done
   273       done
   271   qed
   274   qed
   272   show ?thes2
   275   show ?thes2
   273     apply (simp add: has_field_derivative_iff del: power_Suc)
   276     apply (simp add: has_field_derivative_iff del: power_Suc)
   274     apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \<open>0 < d\<close> ])
   277     apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \<open>0 < d\<close> ])
   340   have *: "\<exists>h. (f' has_field_derivative h) (at z)" if "z \<in> S" for z
   343   have *: "\<exists>h. (f' has_field_derivative h) (at z)" if "z \<in> S" for z
   341   proof -
   344   proof -
   342     obtain r where "r > 0" and r: "cball z r \<subseteq> S"
   345     obtain r where "r > 0" and r: "cball z r \<subseteq> S"
   343       using open_contains_cball \<open>z \<in> S\<close> \<open>open S\<close> by blast
   346       using open_contains_cball \<open>z \<in> S\<close> \<open>open S\<close> by blast
   344     then have holf_cball: "f holomorphic_on cball z r"
   347     then have holf_cball: "f holomorphic_on cball z r"
   345       apply (simp add: holomorphic_on_def)
   348       unfolding holomorphic_on_def
   346       using field_differentiable_at_within field_differentiable_def fder by blast
   349       using field_differentiable_at_within field_differentiable_def fder by fastforce
   347     then have "continuous_on (path_image (circlepath z r)) f"
   350     then have "continuous_on (path_image (circlepath z r)) f"
   348       using \<open>r > 0\<close> by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on])
   351       using \<open>r > 0\<close> by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on])
   349     then have contfpi: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1/(2 * of_real pi*\<i>) * f x)"
   352     then have contfpi: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1/(2 * of_real pi*\<i>) * f x)"
   350       by (auto intro: continuous_intros)+
   353       by (auto intro: continuous_intros)+
   351     have contf_cball: "continuous_on (cball z r) f" using holf_cball
   354     have contf_cball: "continuous_on (cball z r) f" using holf_cball
   370         by (simp add: algebra_simps)
   373         by (simp add: algebra_simps)
   371       then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)"
   374       then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)"
   372         by (simp add: f'_eq)
   375         by (simp add: f'_eq)
   373     } note * = this
   376     } note * = this
   374     show ?thesis
   377     show ?thesis
   375       apply (rule exI)
   378       using Cauchy_next_derivative_circlepath [OF contfpi, of 2 f'] \<open>0 < r\<close> *
   376       apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified])
   379       using centre_in_ball mem_ball by force
   377       apply (simp_all add: \<open>0 < r\<close> * dist_norm)
       
   378       done
       
   379   qed
   380   qed
   380   show ?thesis
   381   show ?thesis
   381     by (simp add: holomorphic_on_open [OF \<open>open S\<close>] *)
   382     by (simp add: holomorphic_on_open [OF \<open>open S\<close>] *)
   382 qed
   383 qed
   383 
   384 
   432                       \<Longrightarrow> contour_integral (linepath a b) f +
   433                       \<Longrightarrow> contour_integral (linepath a b) f +
   433                           contour_integral (linepath b c) f +
   434                           contour_integral (linepath b c) f +
   434                           contour_integral (linepath c a) f = 0"
   435                           contour_integral (linepath c a) f = 0"
   435       by blast
   436       by blast
   436     have az: "dist a z < e" using mem_ball z by blast
   437     have az: "dist a z < e" using mem_ball z by blast
   437     have sb_ball: "ball z (e - dist a z) \<subseteq> ball a e"
       
   438       by (simp add: dist_commute ball_subset_ball_iff)
       
   439     have "\<exists>e>0. f holomorphic_on ball z e"
   438     have "\<exists>e>0. f holomorphic_on ball z e"
   440     proof (intro exI conjI)
   439     proof (intro exI conjI)
   441       have sub_ball: "\<And>y. dist a y < e \<Longrightarrow> closed_segment a y \<subseteq> ball a e"
       
   442         by (meson \<open>0 < e\<close> centre_in_ball convex_ball convex_contains_segment mem_ball)
       
   443       show "f holomorphic_on ball z (e - dist a z)"
   440       show "f holomorphic_on ball z (e - dist a z)"
   444         apply (rule holomorphic_on_subset [OF _ sb_ball])
   441       proof (rule holomorphic_on_subset)
   445         apply (rule derivative_is_holomorphic[OF open_ball])
   442         show "ball z (e - dist a z) \<subseteq> ball a e"
   446         apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a])
   443           by (simp add: dist_commute ball_subset_ball_iff)
   447            apply (simp_all add: 0 \<open>0 < e\<close> sub_ball)
   444         have sub_ball: "\<And>y. dist a y < e \<Longrightarrow> closed_segment a y \<subseteq> ball a e"
   448         done
   445           by (meson \<open>0 < e\<close> centre_in_ball convex_ball convex_contains_segment mem_ball)
       
   446         show "f holomorphic_on ball a e"
       
   447           using triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]
       
   448             derivative_is_holomorphic[OF open_ball]
       
   449           by (force simp add: 0 \<open>0 < e\<close> sub_ball)
       
   450       qed
   449     qed (simp add: az)
   451     qed (simp add: az)
   450   }
   452   }
   451   then show ?thesis
   453   then show ?thesis
   452     by (simp add: analytic_on_def)
   454     by (simp add: analytic_on_def)
   453 qed
   455 qed
   505 lemma higher_deriv_const [simp]: "(deriv ^^ n) (\<lambda>w. c) = (\<lambda>w. if n=0 then c else 0)"
   507 lemma higher_deriv_const [simp]: "(deriv ^^ n) (\<lambda>w. c) = (\<lambda>w. if n=0 then c else 0)"
   506   by (induction n) auto
   508   by (induction n) auto
   507 
   509 
   508 lemma higher_deriv_ident [simp]:
   510 lemma higher_deriv_ident [simp]:
   509      "(deriv ^^ n) (\<lambda>w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)"
   511      "(deriv ^^ n) (\<lambda>w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)"
   510   apply (induction n, simp)
   512 proof (induction n)
   511   apply (metis higher_deriv_linear lambda_one)
   513   case (Suc n)
   512   done
   514   then show ?case by (metis higher_deriv_linear lambda_one)
       
   515 qed auto
   513 
   516 
   514 lemma higher_deriv_id [simp]:
   517 lemma higher_deriv_id [simp]:
   515      "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)"
   518      "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)"
   516   by (simp add: id_def)
   519   by (simp add: id_def)
   517 
   520 
   518 lemma has_complex_derivative_funpow_1:
   521 lemma has_complex_derivative_funpow_1:
   519      "\<lbrakk>(f has_field_derivative 1) (at z); f z = z\<rbrakk> \<Longrightarrow> (f^^n has_field_derivative 1) (at z)"
   522      "\<lbrakk>(f has_field_derivative 1) (at z); f z = z\<rbrakk> \<Longrightarrow> (f^^n has_field_derivative 1) (at z)"
   520   apply (induction n, auto)
   523 proof (induction n)
   521   apply (simp add: id_def)
   524   case 0
   522   by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral)
   525   then show ?case
       
   526     by (simp add: id_def)
       
   527 next
       
   528   case (Suc n)
       
   529   then show ?case
       
   530     by (metis DERIV_chain funpow_Suc_right mult.right_neutral)
       
   531 qed
   523 
   532 
   524 lemma higher_deriv_uminus:
   533 lemma higher_deriv_uminus:
   525   assumes "f holomorphic_on S" "open S" and z: "z \<in> S"
   534   assumes "f holomorphic_on S" "open S" and z: "z \<in> S"
   526     shows "(deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)"
   535     shows "(deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)"
   527 using z
   536 using z
   529   case 0 then show ?case by simp
   538   case 0 then show ?case by simp
   530 next
   539 next
   531   case (Suc n z)
   540   case (Suc n z)
   532   have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
   541   have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
   533     using Suc.prems assms has_field_derivative_higher_deriv by auto
   542     using Suc.prems assms has_field_derivative_higher_deriv by auto
   534   have "((deriv ^^ n) (\<lambda>w. - f w) has_field_derivative - deriv ((deriv ^^ n) f) z) (at z)"
   543   have "\<And>x. x \<in> S \<Longrightarrow> - (deriv ^^ n) f x = (deriv ^^ n) (\<lambda>w. - f w) x"
   535     apply (rule has_field_derivative_transform_within_open [of "\<lambda>w. -((deriv ^^ n) f w)"])
   544     by (auto simp add: Suc)
   536        apply (rule derivative_eq_intros | rule * refl assms)+
   545   then have "((deriv ^^ n) (\<lambda>w. - f w) has_field_derivative - deriv ((deriv ^^ n) f) z) (at z)"
   537      apply (auto simp add: Suc)
   546     using  has_field_derivative_transform_within_open [of "\<lambda>w. -((deriv ^^ n) f w)"]
   538     done
   547     using "*" DERIV_minus Suc.prems \<open>open S\<close> by blast
   539   then show ?case
   548   then show ?case
   540     by (simp add: DERIV_imp_deriv)
   549     by (simp add: DERIV_imp_deriv)
   541 qed
   550 qed
   542 
   551 
   543 lemma higher_deriv_add:
   552 lemma higher_deriv_add:
   550 next
   559 next
   551   case (Suc n z)
   560   case (Suc n z)
   552   have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
   561   have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
   553           "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)"
   562           "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)"
   554     using Suc.prems assms has_field_derivative_higher_deriv by auto
   563     using Suc.prems assms has_field_derivative_higher_deriv by auto
   555   have "((deriv ^^ n) (\<lambda>w. f w + g w) has_field_derivative
   564   have "\<And>x. x \<in> S \<Longrightarrow> (deriv ^^ n) f x + (deriv ^^ n) g x = (deriv ^^ n) (\<lambda>w. f w + g w) x"
       
   565     by (auto simp add: Suc)
       
   566   then have "((deriv ^^ n) (\<lambda>w. f w + g w) has_field_derivative
   556         deriv ((deriv ^^ n) f) z + deriv ((deriv ^^ n) g) z) (at z)"
   567         deriv ((deriv ^^ n) f) z + deriv ((deriv ^^ n) g) z) (at z)"
   557     apply (rule has_field_derivative_transform_within_open [of "\<lambda>w. (deriv ^^ n) f w + (deriv ^^ n) g w"])
   568     using  has_field_derivative_transform_within_open [of "\<lambda>w. (deriv ^^ n) f w + (deriv ^^ n) g w"]
   558        apply (rule derivative_eq_intros | rule * refl assms)+
   569     using "*" Deriv.field_differentiable_add Suc.prems \<open>open S\<close> by blast
   559      apply (auto simp add: Suc)
       
   560     done
       
   561   then show ?case
   570   then show ?case
   562     by (simp add: DERIV_imp_deriv)
   571     by (simp add: DERIV_imp_deriv)
   563 qed
   572 qed
   564 
   573 
   565 lemma higher_deriv_diff:
   574 lemma higher_deriv_diff:
   566   fixes z::complex
   575   fixes z::complex
   567   assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \<in> S"
   576   assumes "f holomorphic_on S" "g holomorphic_on S" "open S" "z \<in> S"
   568     shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z"
   577     shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z"
   569   apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add)
   578   unfolding diff_conv_add_uminus higher_deriv_add
   570   apply (subst higher_deriv_add)
   579   using assms higher_deriv_add higher_deriv_uminus holomorphic_on_minus by presburger
   571   using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus)
       
   572   done
       
   573 
   580 
   574 lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))"
   581 lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))"
   575   by (cases k) simp_all
   582   by (cases k) simp_all
   576 
   583 
   577 lemma higher_deriv_mult:
   584 lemma higher_deriv_mult:
   596     done
   603     done
   597   have "((deriv ^^ n) (\<lambda>w. f w * g w) has_field_derivative
   604   have "((deriv ^^ n) (\<lambda>w. f w * g w) has_field_derivative
   598          (\<Sum>i = 0..Suc n. (Suc n choose i) * (deriv ^^ i) f z * (deriv ^^ (Suc n - i)) g z))
   605          (\<Sum>i = 0..Suc n. (Suc n choose i) * (deriv ^^ i) f z * (deriv ^^ (Suc n - i)) g z))
   599         (at z)"
   606         (at z)"
   600     apply (rule has_field_derivative_transform_within_open
   607     apply (rule has_field_derivative_transform_within_open
   601         [of "\<lambda>w. (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"])
   608         [of "\<lambda>w. (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)" _ _ S])
   602        apply (simp add: algebra_simps)
   609        apply (simp add: algebra_simps)
   603        apply (rule DERIV_cong [OF DERIV_sum])
   610        apply (rule derivative_eq_intros | simp)+
   604         apply (rule DERIV_cmult)
   611            apply (auto intro: DERIV_mult * \<open>open S\<close> Suc.prems Suc.IH [symmetric])
   605         apply (auto intro: DERIV_mult * sumeq \<open>open S\<close> Suc.prems Suc.IH [symmetric])
   612     by (metis (no_types, lifting) mult.commute sum.cong sumeq)
   606     done
       
   607   then show ?case
   613   then show ?case
   608     unfolding funpow.simps o_apply
   614     unfolding funpow.simps o_apply
   609     by (simp add: DERIV_imp_deriv)
   615     by (simp add: DERIV_imp_deriv)
   610 qed
   616 qed
   611 
   617 
   632     by (meson fg f holomorphic_on_subset image_subset_iff)
   638     by (meson fg f holomorphic_on_subset image_subset_iff)
   633   have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S"
   639   have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S"
   634     by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T)
   640     by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T)
   635   have holo3: "(\<lambda>z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S"
   641   have holo3: "(\<lambda>z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S"
   636     by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros)
   642     by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros)
   637   have holo1: "(\<lambda>w. f (u * w)) holomorphic_on S"
   643   have "(*) u holomorphic_on S" "f holomorphic_on (*) u ` S"
   638     apply (rule holomorphic_on_compose [where g=f, unfolded o_def])
   644     by (rule holo0 holomorphic_intros)+
   639     apply (rule holo0 holomorphic_intros)+
   645   then have holo1: "(\<lambda>w. f (u * w)) holomorphic_on S"
   640     done
   646     by (rule holomorphic_on_compose [where g=f, unfolded o_def])
   641   have "deriv ((deriv ^^ n) (\<lambda>w. f (u * w))) z = deriv (\<lambda>z. u^n * (deriv ^^ n) f (u*z)) z"
   647   have "deriv ((deriv ^^ n) (\<lambda>w. f (u * w))) z = deriv (\<lambda>z. u^n * (deriv ^^ n) f (u*z)) z"
   642     apply (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems])
   648   proof (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems])
   643     apply (rule holomorphic_higher_deriv [OF holo1 S])
   649     show "(deriv ^^ n) (\<lambda>w. f (u * w)) holomorphic_on S"
   644     apply (simp add: Suc.IH)
   650       by (rule holomorphic_higher_deriv [OF holo1 S])
   645     done
   651   qed (simp add: Suc.IH)
   646   also have "\<dots> = u^n * deriv (\<lambda>z. (deriv ^^ n) f (u * z)) z"
   652   also have "\<dots> = u^n * deriv (\<lambda>z. (deriv ^^ n) f (u * z)) z"
   647     apply (rule deriv_cmult)
   653   proof -
   648     apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems])
   654     have "(deriv ^^ n) f analytic_on T"
   649     apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and T=T, unfolded o_def])
   655       by (simp add: analytic_on_open f holomorphic_higher_deriv T)
   650       apply (simp)
   656     then have "(\<lambda>w. (deriv ^^ n) f (u * w)) analytic_on S"
   651      apply (simp add: analytic_on_open f holomorphic_higher_deriv T)
   657     proof -
   652     apply (blast intro: fg)
   658       have "(deriv ^^ n) f \<circ> (*) u holomorphic_on S"
   653     done
   659         by (simp add: holo2 holomorphic_on_compose)
       
   660       then show ?thesis
       
   661         by (simp add: S analytic_on_open o_def)
       
   662     qed
       
   663     then show ?thesis
       
   664       by (intro deriv_cmult analytic_on_imp_differentiable_at [OF _ Suc.prems])
       
   665   qed
   654   also have "\<dots> = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)"
   666   also have "\<dots> = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)"
   655       apply (subst deriv_chain [where g = "(deriv ^^ n) f" and f = "(*) u", unfolded o_def])
   667   proof -
   656       apply (rule derivative_intros)
   668     have "(deriv ^^ n) f field_differentiable at (u * z)"
   657       using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv T apply blast
   669       using Suc.prems T f fg holomorphic_higher_deriv holomorphic_on_imp_differentiable_at by blast
   658       apply (simp)
   670     then show ?thesis
   659       done
   671       by (simp add: deriv_compose_linear)
       
   672   qed
   660   finally show ?case
   673   finally show ?case
   661     by simp
   674     by simp
   662 qed
   675 qed
   663 
   676 
   664 lemma higher_deriv_add_at:
   677 lemma higher_deriv_add_at:
   894       using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \<open>0 < e\<close> \<open>0 < B\<close> k by force
   907       using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \<open>0 < e\<close> \<open>0 < B\<close> k by force
   895     have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) - f u / (u - w)) < e"
   908     have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) - f u / (u - w)) < e"
   896          if "n \<le> N" and r: "r = dist z u"  for N u
   909          if "n \<le> N" and r: "r = dist z u"  for N u
   897     proof -
   910     proof -
   898       have N: "((r - k) / r) ^ N < e / B * k"
   911       have N: "((r - k) / r) ^ N < e / B * k"
   899         apply (rule le_less_trans [OF power_decreasing n])
   912         using le_less_trans [OF power_decreasing n]
   900         using  \<open>n \<le> N\<close> k by auto
   913         using \<open>n \<le> N\<close> k by auto
   901       have u [simp]: "(u \<noteq> z) \<and> (u \<noteq> w)"
   914       have u [simp]: "(u \<noteq> z) \<and> (u \<noteq> w)"
   902         using \<open>0 < r\<close> r w by auto
   915         using \<open>0 < r\<close> r w by auto
   903       have wzu_not1: "(w - z) / (u - z) \<noteq> 1"
   916       have wzu_not1: "(w - z) / (u - z) \<noteq> 1"
   904         by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w)
   917         by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w)
   905       have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) * (u - w) - f u)
   918       have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) * (u - w) - f u)
   916         by (simp add: algebra_simps)
   929         by (simp add: algebra_simps)
   917       also have "\<dots> = norm (w - z) ^ N * norm (f u) / r ^ N"
   930       also have "\<dots> = norm (w - z) ^ N * norm (f u) / r ^ N"
   918         by (simp add: norm_mult norm_power norm_minus_commute)
   931         by (simp add: norm_mult norm_power norm_minus_commute)
   919       also have "\<dots> \<le> (((r - k)/r)^N) * B"
   932       also have "\<dots> \<le> (((r - k)/r)^N) * B"
   920         using \<open>0 < r\<close> w k
   933         using \<open>0 < r\<close> w k
   921         apply (simp add: divide_simps)
   934         by (simp add: B divide_simps mult_mono r wz_eq)
   922         apply (rule mult_mono [OF power_mono])
       
   923         apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r)
       
   924         done
       
   925       also have "\<dots> < e * k"
   935       also have "\<dots> < e * k"
   926         using \<open>0 < B\<close> N by (simp add: divide_simps)
   936         using \<open>0 < B\<close> N by (simp add: divide_simps)
   927       also have "\<dots> \<le> e * norm (u - w)"
   937       also have "\<dots> \<le> e * norm (u - w)"
   928         using r kle \<open>0 < e\<close> by (simp add: dist_commute dist_norm)
   938         using r kle \<open>0 < e\<close> by (simp add: dist_commute dist_norm)
   929       finally show ?thesis
   939       finally show ?thesis
   931     qed
   941     qed
   932     with \<open>0 < r\<close> show "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>sphere z r.
   942     with \<open>0 < r\<close> show "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>sphere z r.
   933                 norm ((\<Sum>k<n. (w - z) ^ k * (f x / (x - z) ^ Suc k)) - f x / (x - w)) < e"
   943                 norm ((\<Sum>k<n. (w - z) ^ k * (f x / (x - z) ^ Suc k)) - f x / (x - w)) < e"
   934       by (auto simp: mult_ac less_imp_le eventually_sequentially Ball_def)
   944       by (auto simp: mult_ac less_imp_le eventually_sequentially Ball_def)
   935   qed
   945   qed
       
   946   have \<section>: "\<And>x k. k\<in> {..<x} \<Longrightarrow>
       
   947            (\<lambda>u. (w - z) ^ k * (f u / (u - z) ^ Suc k)) contour_integrable_on circlepath z r"
       
   948     using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] by (simp add: field_simps)
   936   have eq: "\<forall>\<^sub>F x in sequentially.
   949   have eq: "\<forall>\<^sub>F x in sequentially.
   937              contour_integral (circlepath z r) (\<lambda>u. \<Sum>k<x. (w - z) ^ k * (f u / (u - z) ^ Suc k)) =
   950              contour_integral (circlepath z r) (\<lambda>u. \<Sum>k<x. (w - z) ^ k * (f u / (u - z) ^ Suc k)) =
   938              (\<Sum>k<x. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z) ^ k)"
   951              (\<Sum>k<x. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z) ^ k)"
   939     apply (rule eventuallyI)
   952     apply (rule eventuallyI)
   940     apply (subst contour_integral_sum, simp)
   953     apply (subst contour_integral_sum, simp)
   941     using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps)
   954     apply (simp_all only: \<section> contour_integral_lmul cint algebra_simps)
   942     apply (simp only: contour_integral_lmul cint algebra_simps)
       
   943     done
   955     done
   944   have cic: "\<And>u. (\<lambda>y. \<Sum>k<u. (w - z) ^ k * (f y / (y - z) ^ Suc k)) contour_integrable_on circlepath z r"
   956   have "\<And>u k. k \<in> {..<u} \<Longrightarrow> (\<lambda>x. f x / (x - z) ^ Suc k) contour_integrable_on circlepath z r"
   945     apply (intro contour_integrable_sum contour_integrable_lmul, simp)
       
   946     using \<open>0 < r\<close> by (force intro!: Cauchy_higher_derivative_integral_circlepath [OF contf holf])
   957     using \<open>0 < r\<close> by (force intro!: Cauchy_higher_derivative_integral_circlepath [OF contf holf])
   947   have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
   958   then have "\<And>u. (\<lambda>y. \<Sum>k<u. (w - z) ^ k * (f y / (y - z) ^ Suc k)) contour_integrable_on circlepath z r"
       
   959     by (intro contour_integrable_sum contour_integrable_lmul, simp)
       
   960   then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
   948         sums contour_integral (circlepath z r) (\<lambda>u. f u/(u - w))"
   961         sums contour_integral (circlepath z r) (\<lambda>u. f u/(u - w))"
   949     unfolding sums_def
   962     unfolding sums_def using \<open>0 < r\<close> 
   950     apply (intro Lim_transform_eventually [OF _ eq] contour_integral_uniform_limit_circlepath [OF eventuallyI ul] cic)
   963     by (intro Lim_transform_eventually [OF _ eq] contour_integral_uniform_limit_circlepath [OF eventuallyI ul]) auto
   951     using \<open>0 < r\<close> apply auto
       
   952     done
       
   953   then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
   964   then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
   954              sums (2 * of_real pi * \<i> * f w)"
   965              sums (2 * of_real pi * \<i> * f w)"
   955     using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]])
   966     using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]])
   956   then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z)^k / (\<i> * (of_real pi * 2)))
   967   then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z)^k / (\<i> * (of_real pi * 2)))
   957             sums ((2 * of_real pi * \<i> * f w) / (\<i> * (complex_of_real pi * 2)))"
   968             sums ((2 * of_real pi * \<i> * f w) / (\<i> * (complex_of_real pi * 2)))"
   982       by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def)
   993       by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def)
   983     then show "0 < 1 + \<bar>B\<bar> + cmod z"
   994     then show "0 < 1 + \<bar>B\<bar> + cmod z"
   984       by linarith
   995       by linarith
   985   qed
   996   qed
   986   have *: "((\<lambda>u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \<i> * f z) (circlepath z R)"
   997   have *: "((\<lambda>u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \<i> * f z) (circlepath z R)"
   987     apply (rule Cauchy_integral_circlepath)
   998     using continuous_on_subset holf  holomorphic_on_subset \<open>0 < R\<close>
   988     using \<open>R > 0\<close> apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+
   999     by (force intro: holomorphic_on_imp_continuous_on Cauchy_integral_circlepath)
   989     done
       
   990   have "cmod (x - z) = R \<Longrightarrow> cmod (f x) * 2 < cmod (f z)" for x
  1000   have "cmod (x - z) = R \<Longrightarrow> cmod (f x) * 2 < cmod (f z)" for x
   991     unfolding R_def
  1001     unfolding R_def
   992     by (rule B) (use norm_triangle_ineq4 [of x z] in auto)
  1002     by (rule B) (use norm_triangle_ineq4 [of x z] in auto)
   993   with \<open>R > 0\<close> fz show False
  1003   with \<open>R > 0\<close> fz show False
   994     using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"]
  1004     using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"]
  1007 proof -
  1017 proof -
  1008   { assume f: "\<And>z. f z \<noteq> 0"
  1018   { assume f: "\<And>z. f z \<noteq> 0"
  1009     have 1: "(\<lambda>x. 1 / f x) holomorphic_on UNIV"
  1019     have 1: "(\<lambda>x. 1 / f x) holomorphic_on UNIV"
  1010       by (simp add: holomorphic_on_divide assms f)
  1020       by (simp add: holomorphic_on_divide assms f)
  1011     have 2: "((\<lambda>x. 1 / f x) \<longlongrightarrow> 0) at_infinity"
  1021     have 2: "((\<lambda>x. 1 / f x) \<longlongrightarrow> 0) at_infinity"
  1012       apply (rule tendstoI [OF eventually_mono])
  1022     proof (rule tendstoI [OF eventually_mono])
  1013       apply (rule_tac B="2/e" in unbounded)
  1023       fix e::real
  1014       apply (simp add: dist_norm norm_divide field_split_simps)
  1024       assume "e > 0"
  1015       done
  1025       show "eventually (\<lambda>x. 2/e \<le> cmod (f x)) at_infinity"
       
  1026         by (rule_tac B="2/e" in unbounded)
       
  1027     qed (simp add: dist_norm norm_divide field_split_simps)
  1016     have False
  1028     have False
  1017       using Liouville_weak_0 [OF 1 2] f by simp
  1029       using Liouville_weak_0 [OF 1 2] f by simp
  1018   }
  1030   }
  1019   then show ?thesis
  1031   then show ?thesis
  1020     using that by blast
  1032     using that by blast
  1068     define d where "d = (r - norm(w - z))"
  1080     define d where "d = (r - norm(w - z))"
  1069     have "0 < d"  "d \<le> r" using w by (auto simp: norm_minus_commute d_def dist_norm)
  1081     have "0 < d"  "d \<le> r" using w by (auto simp: norm_minus_commute d_def dist_norm)
  1070     have dle: "\<And>u. cmod (z - u) = r \<Longrightarrow> d \<le> cmod (u - w)"
  1082     have dle: "\<And>u. cmod (z - u) = r \<Longrightarrow> d \<le> cmod (u - w)"
  1071       unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
  1083       unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
  1072     have ev_int: "\<forall>\<^sub>F n in F. (\<lambda>u. f n u / (u - w)) contour_integrable_on circlepath z r"
  1084     have ev_int: "\<forall>\<^sub>F n in F. (\<lambda>u. f n u / (u - w)) contour_integrable_on circlepath z r"
  1073       apply (rule eventually_mono [OF cont])
       
  1074       using w
  1085       using w
  1075       apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified])
  1086       by (auto intro: eventually_mono [OF cont] Cauchy_higher_derivative_integral_circlepath [where k=0, simplified])
  1076       done
  1087     have "\<And>e. \<lbrakk>0 < r; 0 < d; 0 < e\<rbrakk>
  1077     have ul_less: "uniform_limit (sphere z r) (\<lambda>n x. f n x / (x - w)) (\<lambda>x. g x / (x - w)) F"
  1088          \<Longrightarrow> \<forall>\<^sub>F n in F.
  1078       using greater \<open>0 < d\<close>
  1089                 \<forall>x\<in>sphere z r.
  1079       apply (clarsimp simp add: uniform_limit_iff dist_norm norm_divide diff_divide_distrib [symmetric] divide_simps)
  1090                    x \<noteq> w \<longrightarrow>
       
  1091                    cmod (f n x - g x) < e * cmod (x - w)"
  1080       apply (rule_tac e1="e * d" in eventually_mono [OF uniform_limitD [OF ulim]])
  1092       apply (rule_tac e1="e * d" in eventually_mono [OF uniform_limitD [OF ulim]])
  1081        apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+
  1093        apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+
  1082       done
  1094       done
       
  1095     then have ul_less: "uniform_limit (sphere z r) (\<lambda>n x. f n x / (x - w)) (\<lambda>x. g x / (x - w)) F"
       
  1096       using greater \<open>0 < d\<close>
       
  1097       by (auto simp add: uniform_limit_iff dist_norm norm_divide diff_divide_distrib [symmetric] divide_simps)
  1083     have g_cint: "(\<lambda>u. g u/(u - w)) contour_integrable_on circlepath z r"
  1098     have g_cint: "(\<lambda>u. g u/(u - w)) contour_integrable_on circlepath z r"
  1084       by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \<open>0 < r\<close>])
  1099       by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \<open>0 < r\<close>])
  1085     have cif_tends_cig: "((\<lambda>n. contour_integral(circlepath z r) (\<lambda>u. f n u / (u - w))) \<longlongrightarrow> contour_integral(circlepath z r) (\<lambda>u. g u/(u - w))) F"
  1100     have cif_tends_cig: "((\<lambda>n. contour_integral(circlepath z r) (\<lambda>u. f n u / (u - w))) \<longlongrightarrow> contour_integral(circlepath z r) (\<lambda>u. g u/(u - w))) F"
  1086       by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \<open>0 < r\<close>])
  1101       by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \<open>0 < r\<close>])
  1087     have f_tends_cig: "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> contour_integral (circlepath z r) (\<lambda>u. g u / (u - w))) F"
  1102     have f_tends_cig: "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> contour_integral (circlepath z r) (\<lambda>u. g u / (u - w))) F"
  1088     proof (rule Lim_transform_eventually)
  1103     proof (rule Lim_transform_eventually)
  1089       show "\<forall>\<^sub>F x in F. contour_integral (circlepath z r) (\<lambda>u. f x u / (u - w))
  1104       show "\<forall>\<^sub>F x in F. contour_integral (circlepath z r) (\<lambda>u. f x u / (u - w))
  1090                      = 2 * of_real pi * \<i> * f x w"
  1105                      = 2 * of_real pi * \<i> * f x w"
  1091         apply (rule eventually_mono [OF cont contour_integral_unique [OF Cauchy_integral_circlepath]])
  1106         using w\<open>0 < d\<close> d_def
  1092         using w\<open>0 < d\<close> d_def by auto
  1107         by (auto intro: eventually_mono [OF cont contour_integral_unique [OF Cauchy_integral_circlepath]])
  1093     qed (auto simp: cif_tends_cig)
  1108     qed (auto simp: cif_tends_cig)
  1094     have "\<And>e. 0 < e \<Longrightarrow> \<forall>\<^sub>F n in F. dist (f n w) (g w) < e"
  1109     have "\<And>e. 0 < e \<Longrightarrow> \<forall>\<^sub>F n in F. dist (f n w) (g w) < e"
  1095       by (rule eventually_mono [OF uniform_limitD [OF ulim]]) (use w in auto)
  1110       by (rule eventually_mono [OF uniform_limitD [OF ulim]]) (use w in auto)
  1096     then have "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> 2 * of_real pi * \<i> * g w) F"
  1111     then have "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> 2 * of_real pi * \<i> * g w) F"
  1097       by (rule tendsto_mult_left [OF tendstoI])
  1112       by (rule tendsto_mult_left [OF tendstoI])
  1164       by (force simp: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont])
  1179       by (force simp: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont])
  1165     have 2: "uniform_limit (sphere z r) (\<lambda>n x. f n x / (x - w)\<^sup>2) (\<lambda>x. g x / (x - w)\<^sup>2) F"
  1180     have 2: "uniform_limit (sphere z r) (\<lambda>n x. f n x / (x - w)\<^sup>2) (\<lambda>x. g x / (x - w)\<^sup>2) F"
  1166       unfolding uniform_limit_iff
  1181       unfolding uniform_limit_iff
  1167     proof clarify
  1182     proof clarify
  1168       fix e::real
  1183       fix e::real
  1169       assume "0 < e"
  1184       assume "e > 0"
  1170       with \<open>r > 0\<close> show "\<forall>\<^sub>F n in F. \<forall>x\<in>sphere z r. dist (f n x / (x - w)\<^sup>2) (g x / (x - w)\<^sup>2) < e"
  1185       with \<open>r > 0\<close> 
  1171         apply (simp add: norm_divide field_split_simps sphere_def dist_norm)
  1186       have "\<forall>\<^sub>F n in F. \<forall>x. x \<noteq> w \<longrightarrow> cmod (z - x) = r \<longrightarrow> cmod (f n x - g x) < e * cmod ((x - w)\<^sup>2)"
  1172         apply (rule eventually_mono [OF uniform_limitD [OF ulim], of "e*d"])
  1187         by (force simp: \<open>0 < d\<close> dist_norm dle intro: less_le_trans eventually_mono [OF uniform_limitD [OF ulim], of "e*d"])
  1173          apply (simp add: \<open>0 < d\<close>)
  1188       with \<open>r > 0\<close> \<open>e > 0\<close> 
  1174         apply (force simp: dist_norm dle intro: less_le_trans)
  1189       show "\<forall>\<^sub>F n in F. \<forall>x\<in>sphere z r. dist (f n x / (x - w)\<^sup>2) (g x / (x - w)\<^sup>2) < e"
  1175         done
  1190         by (simp add: norm_divide field_split_simps sphere_def dist_norm)
  1176     qed
  1191     qed
  1177     have "((\<lambda>n. contour_integral (circlepath z r) (\<lambda>x. f n x / (x - w)\<^sup>2))
  1192     have "((\<lambda>n. contour_integral (circlepath z r) (\<lambda>x. f n x / (x - w)\<^sup>2))
  1178              \<longlongrightarrow> contour_integral (circlepath z r) ((\<lambda>x. g x / (x - w)\<^sup>2))) F"
  1193              \<longlongrightarrow> contour_integral (circlepath z r) ((\<lambda>x. g x / (x - w)\<^sup>2))) F"
  1179       by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \<open>0 < r\<close>])
  1194       by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \<open>0 < r\<close>])
  1180     then have tendsto_0: "((\<lambda>n. 1 / (2 * of_real pi * \<i>) * (?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2))) \<longlongrightarrow> 0) F"
  1195     then have tendsto_0: "((\<lambda>n. 1 / (2 * of_real pi * \<i>) * (?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2))) \<longlongrightarrow> 0) F"
  1211         using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r by blast
  1226         using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r by blast
  1212       show "f x holomorphic_on ball z r" for x
  1227       show "f x holomorphic_on ball z r" for x
  1213         by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r)
  1228         by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r)
  1214     qed
  1229     qed
  1215     show ?thesis
  1230     show ?thesis
  1216       apply (rule holomorphic_uniform_limit [OF *])
       
  1217       using \<open>0 < r\<close> centre_in_ball ul
  1231       using \<open>0 < r\<close> centre_in_ball ul
  1218       apply (auto simp: holomorphic_on_open)
  1232       by (auto simp: holomorphic_on_open intro: holomorphic_uniform_limit [OF *])
  1219       done
       
  1220   qed
  1233   qed
  1221   with S show ?thesis
  1234   with S show ?thesis
  1222     by (simp add: holomorphic_on_open)
  1235     by (simp add: holomorphic_on_open)
  1223 qed
  1236 qed
  1224 
  1237 
  1361              ((\<lambda>n. a n * z^n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * z^(n - 1)) sums g' z) \<and> (g has_field_derivative g' z) (at z)"
  1374              ((\<lambda>n. a n * z^n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * z^(n - 1)) sums g' z) \<and> (g has_field_derivative g' z) (at z)"
  1362 proof (cases "0 < r")
  1375 proof (cases "0 < r")
  1363   case True
  1376   case True
  1364     have der: "\<And>n z. ((\<lambda>x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)"
  1377     have der: "\<And>n z. ((\<lambda>x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)"
  1365       by (rule derivative_eq_intros | simp)+
  1378       by (rule derivative_eq_intros | simp)+
  1366     have y_le: "\<lbrakk>cmod (z - y) * 2 < r - cmod z\<rbrakk> \<Longrightarrow> cmod y \<le> cmod (of_real r + of_real (cmod z)) / 2" for z y
  1379     have y_le: "cmod y \<le> cmod (of_real r + of_real (cmod z)) / 2" 
  1367       using \<open>r > 0\<close>
  1380       if "cmod (z - y) * 2 < r - cmod z" for z y
  1368       apply (auto simp: algebra_simps norm_mult norm_divide norm_power simp flip: of_real_add)
  1381     proof -
  1369       using norm_triangle_ineq2 [of y z]
  1382       have "cmod y * 2 \<le> r + cmod z"
  1370       apply (simp only: diff_le_eq norm_minus_commute mult_2)
  1383         using norm_triangle_ineq2 [of y z] that
  1371       done
  1384         by (simp only: diff_le_eq norm_minus_commute mult_2)
       
  1385       then show ?thesis
       
  1386         using \<open>r > 0\<close> 
       
  1387         by (auto simp: algebra_simps norm_mult norm_divide norm_power simp flip: of_real_add)
       
  1388     qed
  1372     have "summable (\<lambda>n. a n * complex_of_real r ^ n)"
  1389     have "summable (\<lambda>n. a n * complex_of_real r ^ n)"
  1373       using assms \<open>r > 0\<close> by simp
  1390       using assms \<open>r > 0\<close> by simp
  1374     moreover have "\<And>z. cmod z < r \<Longrightarrow> cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)"
  1391     moreover have "\<And>z. cmod z < r \<Longrightarrow> cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)"
  1375       using \<open>r > 0\<close>
  1392       using \<open>r > 0\<close>
  1376       by (simp flip: of_real_add)
  1393       by (simp flip: of_real_add)
  1386       done
  1403       done
  1387   then show ?thesis
  1404   then show ?thesis
  1388     by (simp add: ball_def)
  1405     by (simp add: ball_def)
  1389 next
  1406 next
  1390   case False then show ?thesis
  1407   case False then show ?thesis
  1391     apply (simp add: not_less)
  1408     unfolding not_less
  1392     using less_le_trans norm_not_less_zero by blast
  1409     using less_le_trans norm_not_less_zero by blast
  1393 qed
  1410 qed
  1394 
  1411 
  1395 proposition\<^marker>\<open>tag unimportant\<close> power_series_and_derivative:
  1412 proposition\<^marker>\<open>tag unimportant\<close> power_series_and_derivative:
  1396   fixes a :: "nat \<Rightarrow> complex" and r::real
  1413   fixes a :: "nat \<Rightarrow> complex" and r::real
  1429     have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u
  1446     have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u
  1430     proof -
  1447     proof -
  1431       have less: "cmod (z - u) * 2 < cmod (z - w) + r"
  1448       have less: "cmod (z - u) * 2 < cmod (z - w) + r"
  1432         using that dist_triangle2 [of z u w]
  1449         using that dist_triangle2 [of z u w]
  1433         by (simp add: dist_norm [symmetric] algebra_simps)
  1450         by (simp add: dist_norm [symmetric] algebra_simps)
  1434       show ?thesis
  1451       have "(\<lambda>n. a n * (u - z) ^ n) sums g u" "(\<lambda>n. a n * (u - z) ^ n) sums f u"
  1435         apply (rule sums_unique2 [of "\<lambda>n. a n*(u - z)^n"])
  1452         using gg' [of u] less w by (auto simp: assms dist_norm)
  1436         using gg' [of u] less w
  1453       then show ?thesis
  1437         apply (auto simp: assms dist_norm)
  1454         by (metis sums_unique2)
  1438         done
       
  1439     qed
  1455     qed
  1440     have "(f has_field_derivative g' w) (at w)"
  1456     have "(f has_field_derivative g' w) (at w)"
  1441       by (rule has_field_derivative_transform_within [where d="(r - norm(z - w))/2"])
  1457       by (rule has_field_derivative_transform_within [where d="(r - norm(z - w))/2"])
  1442       (use w gg' [of w] in \<open>(force simp: dist_norm)+\<close>)
  1458       (use w gg' [of w] in \<open>(force simp: dist_norm)+\<close>)
  1443     then show ?thesis ..
  1459     then show ?thesis ..
  1466 lemma holomorphic_fun_eq_on_ball:
  1482 lemma holomorphic_fun_eq_on_ball:
  1467    "\<lbrakk>f holomorphic_on ball z r; g holomorphic_on ball z r;
  1483    "\<lbrakk>f holomorphic_on ball z r; g holomorphic_on ball z r;
  1468      w \<in> ball z r;
  1484      w \<in> ball z r;
  1469      \<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z\<rbrakk>
  1485      \<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z\<rbrakk>
  1470      \<Longrightarrow> f w = g w"
  1486      \<Longrightarrow> f w = g w"
  1471   apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
  1487   by (auto simp: holomorphic_iff_power_series sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
  1472   apply (auto simp: holomorphic_iff_power_series)
       
  1473   done
       
  1474 
  1488 
  1475 lemma holomorphic_fun_eq_0_on_ball:
  1489 lemma holomorphic_fun_eq_0_on_ball:
  1476    "\<lbrakk>f holomorphic_on ball z r;  w \<in> ball z r;
  1490    "\<lbrakk>f holomorphic_on ball z r;  w \<in> ball z r;
  1477      \<And>n. (deriv ^^ n) f z = 0\<rbrakk>
  1491      \<And>n. (deriv ^^ n) f z = 0\<rbrakk>
  1478      \<Longrightarrow> f w = 0"
  1492      \<Longrightarrow> f w = 0"
  1479   apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
  1493   by (auto simp: holomorphic_iff_power_series sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
  1480   apply (auto simp: holomorphic_iff_power_series)
       
  1481   done
       
  1482 
  1494 
  1483 lemma holomorphic_fun_eq_0_on_connected:
  1495 lemma holomorphic_fun_eq_0_on_connected:
  1484   assumes holf: "f holomorphic_on S" and "open S"
  1496   assumes holf: "f holomorphic_on S" and "open S"
  1485       and cons: "connected S"
  1497       and cons: "connected S"
  1486       and der: "\<And>n. (deriv ^^ n) f z = 0"
  1498       and der: "\<And>n. (deriv ^^ n) f z = 0"
  1488     shows "f w = 0"
  1500     shows "f w = 0"
  1489 proof -
  1501 proof -
  1490   have *: "ball x e \<subseteq> (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
  1502   have *: "ball x e \<subseteq> (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
  1491     if "\<forall>u. (deriv ^^ u) f x = 0" "ball x e \<subseteq> S" for x e
  1503     if "\<forall>u. (deriv ^^ u) f x = 0" "ball x e \<subseteq> S" for x e
  1492   proof -
  1504   proof -
  1493     have "\<And>x' n. dist x x' < e \<Longrightarrow> (deriv ^^ n) f x' = 0"
  1505     have "(deriv ^^ m) ((deriv ^^ n) f) x = 0" for m n
  1494       apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv])
  1506       by (metis funpow_add o_apply that(1))
  1495          apply (rule holomorphic_on_subset [OF holf])
  1507     then have "\<And>x' n. dist x x' < e \<Longrightarrow> (deriv ^^ n) f x' = 0"
  1496       using that apply simp_all
  1508       using \<open>open S\<close> 
  1497       by (metis funpow_add o_apply)
  1509       by (meson holf holomorphic_fun_eq_0_on_ball holomorphic_higher_deriv holomorphic_on_subset mem_ball that(2))
  1498     with that show ?thesis by auto
  1510     with that show ?thesis by auto
  1499   qed
  1511   qed
  1500   have 1: "openin (top_of_set S) (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
  1512   obtain e where "e>0" and e: "ball w e \<subseteq> S" using openE [OF \<open>open S\<close> \<open>w \<in> S\<close>] .
  1501     apply (rule open_subset, force)
  1513   then have holfb: "f holomorphic_on ball w e"
       
  1514     using holf holomorphic_on_subset by blast
       
  1515   have "open (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
  1502     using \<open>open S\<close>
  1516     using \<open>open S\<close>
  1503     apply (simp add: open_contains_ball Ball_def)
  1517     apply (simp add: open_contains_ball Ball_def)
  1504     apply (erule all_forward)
  1518     apply (erule all_forward)
  1505     using "*" by auto blast+
  1519     using "*" by auto blast+
  1506   have 2: "closedin (top_of_set S) (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
  1520   then have "openin (top_of_set S) (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
       
  1521     by (force intro: open_subset)
       
  1522   moreover have "closedin (top_of_set S) (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
  1507     using assms
  1523     using assms
  1508     by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv)
  1524     by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv)
  1509   obtain e where "e>0" and e: "ball w e \<subseteq> S" using openE [OF \<open>open S\<close> \<open>w \<in> S\<close>] .
  1525   moreover have "(\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0}) = S \<Longrightarrow> f w = 0"
  1510   then have holfb: "f holomorphic_on ball w e"
       
  1511     using holf holomorphic_on_subset by blast
       
  1512   have 3: "(\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0}) = S \<Longrightarrow> f w = 0"
       
  1513     using \<open>e>0\<close> e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb])
  1526     using \<open>e>0\<close> e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb])
  1514   show ?thesis
  1527   ultimately show ?thesis
  1515     using cons der \<open>z \<in> S\<close>
  1528     using cons der \<open>z \<in> S\<close>
  1516     apply (simp add: connected_clopen)
  1529     by (auto simp add: connected_clopen)
  1517     apply (drule_tac x="\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0}" in spec)
       
  1518     apply (auto simp: 1 2 3)
       
  1519     done
       
  1520 qed
  1530 qed
  1521 
  1531 
  1522 lemma holomorphic_fun_eq_on_connected:
  1532 lemma holomorphic_fun_eq_on_connected:
  1523   assumes "f holomorphic_on S" "g holomorphic_on S" and "open S"  "connected S"
  1533   assumes "f holomorphic_on S" "g holomorphic_on S" and "open S"  "connected S"
  1524       and "\<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z"
  1534       and "\<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z"
  1564   have F2: "?F field_differentiable at a" if "0 < e" "ball a e \<subseteq> S" for e
  1574   have F2: "?F field_differentiable at a" if "0 < e" "ball a e \<subseteq> S" for e
  1565   proof -
  1575   proof -
  1566     have holfb: "f holomorphic_on ball a e"
  1576     have holfb: "f holomorphic_on ball a e"
  1567       by (rule holomorphic_on_subset [OF holf \<open>ball a e \<subseteq> S\<close>])
  1577       by (rule holomorphic_on_subset [OF holf \<open>ball a e \<subseteq> S\<close>])
  1568     have 2: "?F holomorphic_on ball a e - {a}"
  1578     have 2: "?F holomorphic_on ball a e - {a}"
  1569       apply (simp add: holomorphic_on_def flip: field_differentiable_def)
       
  1570       using mem_ball that
  1579       using mem_ball that
  1571       apply (auto intro: F1 field_differentiable_within_subset)
  1580       by (auto simp add: holomorphic_on_def simp flip: field_differentiable_def intro: F1 field_differentiable_within_subset)
  1572       done
       
  1573     have "isCont (\<lambda>z. if z = a then deriv f a else (f z - f a) / (z - a)) x"
  1581     have "isCont (\<lambda>z. if z = a then deriv f a else (f z - f a) / (z - a)) x"
  1574             if "dist a x < e" for x
  1582             if "dist a x < e" for x
  1575     proof (cases "x=a")
  1583     proof (cases "x=a")
  1576       case True
  1584       case True
  1577       then have "f field_differentiable at a"
  1585       then have "f field_differentiable at a"
  1733         by (simp add: compact_Times)
  1741         by (simp add: compact_Times)
  1734     qed
  1742     qed
  1735     then obtain \<eta> where "\<eta>>0"
  1743     then obtain \<eta> where "\<eta>>0"
  1736         and \<eta>: "\<And>x x'. \<lbrakk>x\<in>?TZ; x'\<in>?TZ; dist x' x < \<eta>\<rbrakk> \<Longrightarrow>
  1744         and \<eta>: "\<And>x x'. \<lbrakk>x\<in>?TZ; x'\<in>?TZ; dist x' x < \<eta>\<rbrakk> \<Longrightarrow>
  1737                          dist ((\<lambda>(x,y). F x y) x') ((\<lambda>(x,y). F x y) x) < \<epsilon>/norm(b - a)"
  1745                          dist ((\<lambda>(x,y). F x y) x') ((\<lambda>(x,y). F x y) x) < \<epsilon>/norm(b - a)"
  1738       apply (rule uniformly_continuous_onE [where e = "\<epsilon>/norm(b - a)"])
  1746       using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close>
  1739       using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> by auto
  1747       by (auto elim: uniformly_continuous_onE [where e = "\<epsilon>/norm(b - a)"])
  1740     have \<eta>: "\<lbrakk>norm (w - x1) \<le> \<delta>;   x2 \<in> closed_segment a b;
  1748     have \<eta>: "\<lbrakk>norm (w - x1) \<le> \<delta>;   x2 \<in> closed_segment a b;
  1741               norm (w - x1') \<le> \<delta>;  x2' \<in> closed_segment a b; norm ((x1', x2') - (x1, x2)) < \<eta>\<rbrakk>
  1749               norm (w - x1') \<le> \<delta>;  x2' \<in> closed_segment a b; norm ((x1', x2') - (x1, x2)) < \<eta>\<rbrakk>
  1742               \<Longrightarrow> norm (F x1' x2' - F x1 x2) \<le> \<epsilon> / cmod (b - a)"
  1750               \<Longrightarrow> norm (F x1' x2' - F x1 x2) \<le> \<epsilon> / cmod (b - a)"
  1743              for x1 x2 x1' x2'
  1751              for x1 x2 x1' x2'
  1744       using \<eta> [of "(x1,x2)" "(x1',x2')"] by (force simp: dist_norm)
  1752       using \<eta> [of "(x1,x2)" "(x1',x2')"] by (force simp: dist_norm)
  1747     proof -
  1755     proof -
  1748       have "(\<lambda>x. F x' x - F w x) contour_integrable_on linepath a b"
  1756       have "(\<lambda>x. F x' x - F w x) contour_integrable_on linepath a b"
  1749         by (simp add: \<open>w \<in> U\<close> cont_dw contour_integrable_diff that)
  1757         by (simp add: \<open>w \<in> U\<close> cont_dw contour_integrable_diff that)
  1750       then have "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>/norm(b - a) * norm(b - a)"
  1758       then have "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>/norm(b - a) * norm(b - a)"
  1751         apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \<eta>])
  1759         apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \<eta>])
  1752         using \<open>0 < \<epsilon>\<close> \<open>0 < \<delta>\<close> that apply (auto simp: norm_minus_commute)
  1760         using \<open>0 < \<epsilon>\<close> \<open>0 < \<delta>\<close> that by (auto simp: norm_minus_commute)
  1753         done
       
  1754       also have "\<dots> = \<epsilon>" using \<open>a \<noteq> b\<close> by simp
  1761       also have "\<dots> = \<epsilon>" using \<open>a \<noteq> b\<close> by simp
  1755       finally show ?thesis .
  1762       finally show ?thesis .
  1756     qed
  1763     qed
  1757     show ?thesis
  1764     show ?thesis
  1758       apply (rule_tac x="min \<delta> \<eta>" in exI)
  1765       apply (rule_tac x="min \<delta> \<eta>" in exI)
  1759       using \<open>0 < \<delta>\<close> \<open>0 < \<eta>\<close>
  1766       using \<open>0 < \<delta>\<close> \<open>0 < \<eta>\<close>
  1760       apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \<open>w \<in> U\<close> intro: le_ee)
  1767       by (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \<open>w \<in> U\<close> intro: le_ee)
  1761       done
       
  1762   qed
  1768   qed
  1763   show ?thesis
  1769   show ?thesis
  1764   proof (cases "a=b")
  1770   proof (cases "a=b")
  1765     case True
       
  1766     then show ?thesis by simp
       
  1767   next
       
  1768     case False
  1771     case False
  1769     show ?thesis
  1772     show ?thesis
  1770       by (rule continuous_onI) (use False in \<open>auto intro: *\<close>)
  1773       by (rule continuous_onI) (use False in \<open>auto intro: *\<close>)
  1771   qed
  1774   qed auto
  1772 qed
  1775 qed
  1773 
  1776 
  1774 text\<open>This version has \<^term>\<open>polynomial_function \<gamma>\<close> as an additional assumption.\<close>
  1777 text\<open>This version has \<^term>\<open>polynomial_function \<gamma>\<close> as an additional assumption.\<close>
  1775 lemma Cauchy_integral_formula_global_weak:
  1778 lemma Cauchy_integral_formula_global_weak:
  1776   assumes "open U" and holf: "f holomorphic_on U"
  1779   assumes "open U" and holf: "f holomorphic_on U"
  1800     have *: "(\<lambda>c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on U"
  1803     have *: "(\<lambda>c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on U"
  1801       by (simp add: holf pole_lemma_open \<open>open U\<close>)
  1804       by (simp add: holf pole_lemma_open \<open>open U\<close>)
  1802     then have "isCont (\<lambda>x. if x = y then deriv f y else (f x - f y) / (x - y)) y"
  1805     then have "isCont (\<lambda>x. if x = y then deriv f y else (f x - f y) / (x - y)) y"
  1803       using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that \<open>open U\<close> by fastforce
  1806       using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that \<open>open U\<close> by fastforce
  1804     then have "continuous_on U (d y)"
  1807     then have "continuous_on U (d y)"
  1805       apply (simp add: d_def continuous_on_eq_continuous_at \<open>open U\<close>, clarify)
  1808       using "*" d_def holomorphic_on_imp_continuous_on by auto
  1806       using * holomorphic_on_def
       
  1807       by (meson field_differentiable_within_open field_differentiable_imp_continuous_at \<open>open U\<close>)
       
  1808     moreover have "d y holomorphic_on U - {y}"
  1809     moreover have "d y holomorphic_on U - {y}"
  1809     proof -
  1810     proof -
  1810       have "\<And>w. w \<in> U - {y} \<Longrightarrow>
  1811       have "(\<lambda>w. if w = y then deriv f y else (f w - f y) / (w - y)) field_differentiable at w"
  1811                  (\<lambda>w. if w = y then deriv f y else (f w - f y) / (w - y)) field_differentiable at w"
  1812         if "w \<in> U - {y}" for w
  1812         apply (rule_tac d="dist w y" and f = "\<lambda>w. (f w - f y)/(w - y)" in field_differentiable_transform_within)
  1813       proof (rule field_differentiable_transform_within)
  1813            apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros)
  1814         show "(\<lambda>w. (f w - f y) / (w - y)) field_differentiable at w"
  1814         using \<open>open U\<close> holf holomorphic_on_imp_differentiable_at by blast
  1815           using that \<open>open U\<close> holf 
       
  1816           by (auto intro!: holomorphic_on_imp_differentiable_at derivative_intros)
       
  1817         show "dist w y > 0"
       
  1818           using that by auto
       
  1819       qed (auto simp: dist_commute)
  1815       then show ?thesis
  1820       then show ?thesis
  1816         unfolding field_differentiable_def by (simp add: d_def holomorphic_on_open \<open>open U\<close> open_delete)
  1821         unfolding field_differentiable_def by (simp add: d_def holomorphic_on_open \<open>open U\<close> open_delete)
  1817     qed
  1822     qed
  1818     ultimately show ?thesis
  1823     ultimately show ?thesis
  1819       by (rule no_isolated_singularity) (auto simp: \<open>open U\<close>)
  1824       by (rule no_isolated_singularity) (auto simp: \<open>open U\<close>)
  1830   have U: "((d z) has_contour_integral h z) \<gamma>" if "z \<in> U" for z
  1835   have U: "((d z) has_contour_integral h z) \<gamma>" if "z \<in> U" for z
  1831   proof -
  1836   proof -
  1832     have "d z holomorphic_on U"
  1837     have "d z holomorphic_on U"
  1833       by (simp add: hol_d that)
  1838       by (simp add: hol_d that)
  1834     with that show ?thesis
  1839     with that show ?thesis
  1835     apply (simp add: h_def)
  1840       by (metis Diff_subset \<open>valid_path \<gamma>\<close> \<open>open U\<close> contour_integrable_holomorphic_simple h_def has_contour_integral_integral pasz subset_trans)
  1836       by (meson Diff_subset \<open>open U\<close> \<open>valid_path \<gamma>\<close> contour_integrable_holomorphic_simple has_contour_integral_integral pasz subset_trans)
       
  1837   qed
  1841   qed
  1838   have V: "((\<lambda>w. f w / (w - z)) has_contour_integral h z) \<gamma>" if z: "z \<in> v" for z
  1842   have V: "((\<lambda>w. f w / (w - z)) has_contour_integral h z) \<gamma>" if z: "z \<in> v" for z
  1839   proof -
  1843   proof -
  1840     have 0: "0 = (f z) * 2 * of_real (2 * pi) * \<i> * winding_number \<gamma> z"
  1844     have 0: "0 = (f z) * 2 * of_real (2 * pi) * \<i> * winding_number \<gamma> z"
  1841       using v_def z by auto
  1845       using v_def z by auto
  1845       using has_contour_integral_lmul by fastforce
  1849       using has_contour_integral_lmul by fastforce
  1846     then have "((\<lambda>x. f z / (x - z)) has_contour_integral 0) \<gamma>"
  1850     then have "((\<lambda>x. f z / (x - z)) has_contour_integral 0) \<gamma>"
  1847       by (simp add: field_split_simps)
  1851       by (simp add: field_split_simps)
  1848     moreover have "((\<lambda>x. (f x - f z) / (x - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
  1852     moreover have "((\<lambda>x. (f x - f z) / (x - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
  1849       using z
  1853       using z
  1850       apply (auto simp: v_def)
  1854       apply (simp add: v_def)
  1851       apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy)
  1855       apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy)
  1852       done
  1856       done
  1853     ultimately have *: "((\<lambda>x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \<gamma> (d z))) \<gamma>"
  1857     ultimately have *: "((\<lambda>x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \<gamma> (d z))) \<gamma>"
  1854       by (rule has_contour_integral_add)
  1858       by (rule has_contour_integral_add)
  1855     have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
  1859     have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
  1856             if  "z \<in> U"
  1860       if "z \<in> U"
  1857       using * by (auto simp: divide_simps has_contour_integral_eq)
  1861       using * by (auto simp: divide_simps has_contour_integral_eq)
  1858     moreover have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (\<lambda>w. f w / (w - z))) \<gamma>"
  1862     moreover have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (\<lambda>w. f w / (w - z))) \<gamma>"
  1859             if "z \<notin> U"
  1863       if "z \<notin> U"
  1860       apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where S=U]])
  1864     proof (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where S=U]])
  1861       using U pasz \<open>valid_path \<gamma>\<close> that
  1865       show "(\<lambda>w. f w / (w - z)) holomorphic_on U"
  1862       apply (auto intro: holomorphic_on_imp_continuous_on hol_d)
  1866         by (rule holomorphic_intros assms | use that in force)+
  1863        apply (rule continuous_intros conf holomorphic_intros holf assms | force)+
  1867     qed (use \<open>open U\<close> pasz \<open>valid_path \<gamma>\<close> in auto)
  1864       done
       
  1865     ultimately show ?thesis
  1868     ultimately show ?thesis
  1866       using z by (simp add: h_def)
  1869       using z by (simp add: h_def)
  1867   qed
  1870   qed
  1868   have znot: "z \<notin> path_image \<gamma>"
  1871   have znot: "z \<notin> path_image \<gamma>"
  1869     using pasz by blast
  1872     using pasz by blast
  1870   obtain d0 where "d0>0" and d0: "\<And>x y. x \<in> path_image \<gamma> \<Longrightarrow> y \<in> - U \<Longrightarrow> d0 \<le> dist x y"
  1873   obtain d0 where "d0>0" and d0: "\<And>x y. x \<in> path_image \<gamma> \<Longrightarrow> y \<in> - U \<Longrightarrow> d0 \<le> dist x y"
  1871     using separate_compact_closed [of "path_image \<gamma>" "-U"] pasz \<open>open U\<close> \<open>path \<gamma>\<close> compact_path_image
  1874     using separate_compact_closed [of "path_image \<gamma>" "-U"] pasz \<open>open U\<close> \<open>path \<gamma>\<close> compact_path_image
  1872     by blast    
  1875     by blast    
  1873   obtain dd where "0 < dd" and dd: "{y + k | y k. y \<in> path_image \<gamma> \<and> k \<in> ball 0 dd} \<subseteq> U"
  1876   obtain dd where "0 < dd" and dd: "{y + k | y k. y \<in> path_image \<gamma> \<and> k \<in> ball 0 dd} \<subseteq> U"
  1874     apply (rule that [of "d0/2"])
  1877   proof
  1875     using \<open>0 < d0\<close>
  1878     show "0 < d0 / 2" using \<open>0 < d0\<close> by auto
  1876     apply (auto simp: dist_norm dest: d0)
  1879   qed (use \<open>0 < d0\<close> d0 in \<open>force simp: dist_norm\<close>)
  1877     done
  1880   define T where "T \<equiv> {y + k |y k. y \<in> path_image \<gamma> \<and> k \<in> cball 0 (dd / 2)}"
  1878   have "\<And>x x'. \<lbrakk>x \<in> path_image \<gamma>; dist x x' * 2 < dd\<rbrakk> \<Longrightarrow> \<exists>y k. x' = y + k \<and> y \<in> path_image \<gamma> \<and> dist 0 k * 2 \<le> dd"
  1881   have "\<And>x x'. \<lbrakk>x \<in> path_image \<gamma>; dist x x' * 2 < dd\<rbrakk> \<Longrightarrow> \<exists>y k. x' = y + k \<and> y \<in> path_image \<gamma> \<and> dist 0 k * 2 \<le> dd"
  1879     apply (rule_tac x=x in exI)
  1882     apply (rule_tac x=x in exI)
  1880     apply (rule_tac x="x'-x" in exI)
  1883     apply (rule_tac x="x'-x" in exI)
  1881     apply (force simp: dist_norm)
  1884     apply (force simp: dist_norm)
  1882     done
  1885     done
  1883   then have 1: "path_image \<gamma> \<subseteq> interior {y + k |y k. y \<in> path_image \<gamma> \<and> k \<in> cball 0 (dd / 2)}"
  1886   then have subt: "path_image \<gamma> \<subseteq> interior T"
  1884     apply (clarsimp simp add: mem_interior)
  1887     using \<open>0 < dd\<close> 
  1885     using \<open>0 < dd\<close>
  1888     apply (clarsimp simp add: mem_interior T_def)
  1886     apply (rule_tac x="dd/2" in exI, auto)
  1889     apply (rule_tac x="dd/2" in exI, auto)
  1887     done
  1890     done
  1888   obtain T where "compact T" and subt: "path_image \<gamma> \<subseteq> interior T" and T: "T \<subseteq> U"
  1891   have "compact T"
  1889     apply (rule that [OF _ 1])
  1892     unfolding T_def
  1890     apply (fastforce simp add: \<open>valid_path \<gamma>\<close> compact_valid_path_image intro!: compact_sums)
  1893     by (fastforce simp add: \<open>valid_path \<gamma>\<close> compact_valid_path_image intro!: compact_sums)
  1891     apply (rule order_trans [OF _ dd])
  1894   have T: "T \<subseteq> U"
  1892     using \<open>0 < dd\<close> by fastforce
  1895     unfolding T_def using \<open>0 < dd\<close> dd by fastforce
  1893   obtain L where "L>0"
  1896   obtain L where "L>0"
  1894            and L: "\<And>f B. \<lbrakk>f holomorphic_on interior T; \<And>z. z\<in>interior T \<Longrightarrow> cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
  1897            and L: "\<And>f B. \<lbrakk>f holomorphic_on interior T; \<And>z. z\<in>interior T \<Longrightarrow> cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
  1895                          cmod (contour_integral \<gamma> f) \<le> L * B"
  1898                          cmod (contour_integral \<gamma> f) \<le> L * B"
  1896       using contour_integral_bound_exists [OF open_interior \<open>valid_path \<gamma>\<close> subt]
  1899       using contour_integral_bound_exists [OF open_interior \<open>valid_path \<gamma>\<close> subt]
  1897       by blast
  1900       by blast
  1907     with le have ybig: "norm y > C" by force
  1910     with le have ybig: "norm y > C" by force
  1908     with C have "y \<notin> T"  by force
  1911     with C have "y \<notin> T"  by force
  1909     then have ynot: "y \<notin> path_image \<gamma>"
  1912     then have ynot: "y \<notin> path_image \<gamma>"
  1910       using subt interior_subset by blast
  1913       using subt interior_subset by blast
  1911     have [simp]: "winding_number \<gamma> y = 0"
  1914     have [simp]: "winding_number \<gamma> y = 0"
  1912       apply (rule winding_number_zero_outside [of _ "cball 0 C"])
  1915     proof (rule winding_number_zero_outside)
  1913       using ybig interior_subset subt
  1916       show "path_image \<gamma> \<subseteq> cball 0 C"
  1914       apply (force simp: loop \<open>path \<gamma>\<close> dist_norm intro!: C)+
  1917         by (meson C interior_subset mem_cball_0 subset_eq subt)
  1915       done
  1918     qed (use ybig loop \<open>path \<gamma>\<close> in auto)
  1916     have [simp]: "h y = contour_integral \<gamma> (\<lambda>w. f w/(w - y))"
  1919     have [simp]: "h y = contour_integral \<gamma> (\<lambda>w. f w/(w - y))"
  1917       by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V)
  1920       by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V)
  1918     have holint: "(\<lambda>w. f w / (w - y)) holomorphic_on interior T"
  1921     have holint: "(\<lambda>w. f w / (w - y)) holomorphic_on interior T"
  1919       apply (rule holomorphic_on_divide)
  1922     proof (intro holomorphic_intros)
  1920       using holf holomorphic_on_subset interior_subset T apply blast
  1923       show "f holomorphic_on interior T"
  1921       apply (rule holomorphic_intros)+
  1924         using holf holomorphic_on_subset interior_subset T by blast
  1922       using \<open>y \<notin> T\<close> interior_subset by auto
  1925     qed (use \<open>y \<notin> T\<close> interior_subset in auto)
  1923     have leD: "cmod (f z / (z - y)) \<le> D * (e / L / D)" if z: "z \<in> interior T" for z
  1926     have leD: "cmod (f z / (z - y)) \<le> D * (e / L / D)" if z: "z \<in> interior T" for z
  1924     proof -
  1927     proof -
  1925       have "D * L / e + cmod z \<le> cmod y"
  1928       have "D * L / e + cmod z \<le> cmod y"
  1926         using le C [of z] z using interior_subset by force
  1929         using le C [of z] z using interior_subset by force
  1927       then have DL2: "D * L / e \<le> cmod (z - y)"
  1930       then have DL2: "D * L / e \<le> cmod (z - y)"
  1928         using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute)
  1931         using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute)
  1929       have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))"
  1932       have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))"
  1930         by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse)
  1933         by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse)
  1931       also have "\<dots> \<le> D * (e / L / D)"
  1934       also have "\<dots> \<le> D * (e / L / D)"
  1932         apply (rule mult_mono)
  1935       proof (rule mult_mono)
  1933         using that D interior_subset apply blast
  1936         show "cmod (f z) \<le> D"
  1934         using \<open>L>0\<close> \<open>e>0\<close> \<open>D>0\<close> DL2
  1937           using D interior_subset z by blast 
  1935         apply (auto simp: norm_divide field_split_simps)
  1938         show "inverse (cmod (z - y)) \<le> e / L / D" "D \<ge> 0"
  1936         done
  1939           using \<open>L>0\<close> \<open>e>0\<close> \<open>D>0\<close> DL2 by (auto simp: norm_divide field_split_simps)
       
  1940       qed auto
  1937       finally show ?thesis .
  1941       finally show ?thesis .
  1938     qed
  1942     qed
  1939     have "dist (h y) 0 = cmod (contour_integral \<gamma> (\<lambda>w. f w / (w - y)))"
  1943     have "dist (h y) 0 = cmod (contour_integral \<gamma> (\<lambda>w. f w / (w - y)))"
  1940       by (simp add: dist_norm)
  1944       by (simp add: dist_norm)
  1941     also have "\<dots> \<le> L * (D * (e / L / D))"
  1945     also have "\<dots> \<le> L * (D * (e / L / D))"
  1955       apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+
  1959       apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+
  1956       done
  1960       done
  1957     have con_fstsnd: "continuous_on UNIV (\<lambda>x. (fst x - snd x) ::complex)"
  1961     have con_fstsnd: "continuous_on UNIV (\<lambda>x. (fst x - snd x) ::complex)"
  1958       by (rule continuous_intros)+
  1962       by (rule continuous_intros)+
  1959     have open_uu_Id: "open (U \<times> U - Id)"
  1963     have open_uu_Id: "open (U \<times> U - Id)"
  1960       apply (rule open_Diff)
  1964     proof (rule open_Diff)
  1961       apply (simp add: open_Times \<open>open U\<close>)
  1965       show "open (U \<times> U)"
  1962       using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0]
  1966         by (simp add: open_Times \<open>open U\<close>)
  1963       apply (auto simp: Id_fstsnd_eq algebra_simps)
  1967       show "closed (Id :: complex rel)"
  1964       done
  1968         using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0]
       
  1969         by (auto simp: Id_fstsnd_eq algebra_simps)
       
  1970     qed
  1965     have con_derf: "continuous (at z) (deriv f)" if "z \<in> U" for z
  1971     have con_derf: "continuous (at z) (deriv f)" if "z \<in> U" for z
  1966       apply (rule continuous_on_interior [of U])
  1972     proof (rule continuous_on_interior)
  1967       apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \<open>open U\<close>)
  1973       show "continuous_on U (deriv f)"
  1968       by (simp add: interior_open that \<open>open U\<close>)
  1974         by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \<open>open U\<close>)
       
  1975     qed (simp add: interior_open that \<open>open U\<close>)
  1969     have tendsto_f': "((\<lambda>(x,y). if y = x then deriv f (x)
  1976     have tendsto_f': "((\<lambda>(x,y). if y = x then deriv f (x)
  1970                                 else (f (y) - f (x)) / (y - x)) \<longlongrightarrow> deriv f x)
  1977                                 else (f (y) - f (x)) / (y - x)) \<longlongrightarrow> deriv f x)
  1971                       (at (x, x) within U \<times> U)" if "x \<in> U" for x
  1978                       (at (x, x) within U \<times> U)" if "x \<in> U" for x
  1972     proof (rule Lim_withinI)
  1979     proof (rule Lim_withinI)
  1973       fix e::real assume "0 < e"
  1980       fix e::real assume "0 < e"
  1979       have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e"
  1986       have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e"
  1980                     if "z' \<noteq> x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2"
  1987                     if "z' \<noteq> x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2"
  1981                  for x' z'
  1988                  for x' z'
  1982       proof -
  1989       proof -
  1983         have cs_less: "w \<in> closed_segment x' z' \<Longrightarrow> cmod (w - x) \<le> norm (x'-x, z'-x)" for w
  1990         have cs_less: "w \<in> closed_segment x' z' \<Longrightarrow> cmod (w - x) \<le> norm (x'-x, z'-x)" for w
  1984           apply (drule segment_furthest_le [where y=x])
  1991           using segment_furthest_le [of w x' z' x]
  1985           by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans)
  1992           by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans)
  1986         have derf_le: "w \<in> closed_segment x' z' \<Longrightarrow> z' \<noteq> x' \<Longrightarrow> cmod (deriv f w - deriv f x) \<le> e" for w
  1993         have derf_le: "w \<in> closed_segment x' z' \<Longrightarrow> z' \<noteq> x' \<Longrightarrow> cmod (deriv f w - deriv f x) \<le> e" for w
  1987           by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans)
  1994           by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans)
  1988         have f_has_der: "\<And>x. x \<in> U \<Longrightarrow> (f has_field_derivative deriv f x) (at x within U)"
  1995         have f_has_der: "\<And>x. x \<in> U \<Longrightarrow> (f has_field_derivative deriv f x) (at x within U)"
  1989           by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \<open>open U\<close>)
  1996           by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \<open>open U\<close>)
  2005       show "\<exists>d>0. \<forall>xa\<in>U \<times> U.
  2012       show "\<exists>d>0. \<forall>xa\<in>U \<times> U.
  2006                   0 < dist xa (x, x) \<and> dist xa (x, x) < d \<longrightarrow>
  2013                   0 < dist xa (x, x) \<and> dist xa (x, x) < d \<longrightarrow>
  2007                   dist (case xa of (x, y) \<Rightarrow> if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \<le> e"
  2014                   dist (case xa of (x, y) \<Rightarrow> if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \<le> e"
  2008         apply (rule_tac x="min k1 k2" in exI)
  2015         apply (rule_tac x="min k1 k2" in exI)
  2009         using \<open>k1>0\<close> \<open>k2>0\<close> \<open>e>0\<close>
  2016         using \<open>k1>0\<close> \<open>k2>0\<close> \<open>e>0\<close>
  2010         apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le)
  2017         by (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le)
  2011         done
       
  2012     qed
  2018     qed
  2013     have con_pa_f: "continuous_on (path_image \<gamma>) f"
  2019     have con_pa_f: "continuous_on (path_image \<gamma>) f"
  2014       by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt T)
  2020       by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt T)
  2015     have le_B: "\<And>T. T \<in> {0..1} \<Longrightarrow> cmod (vector_derivative \<gamma> (at T)) \<le> B"
  2021     have le_B: "\<And>T. T \<in> {0..1} \<Longrightarrow> cmod (vector_derivative \<gamma> (at T)) \<le> B"
  2016       apply (rule B)
  2022       using \<gamma>' B by (simp add: path_image_def vector_derivative_at rev_image_eqI)
  2017       using \<gamma>' using path_image_def vector_derivative_at by fastforce
       
  2018     have f_has_cint: "\<And>w. w \<in> v - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f u / (u - w) ^ 1) has_contour_integral h w) \<gamma>"
  2023     have f_has_cint: "\<And>w. w \<in> v - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f u / (u - w) ^ 1) has_contour_integral h w) \<gamma>"
  2019       by (simp add: V)
  2024       by (simp add: V)
  2020     have cond_uu: "continuous_on (U \<times> U) (\<lambda>(x,y). d x y)"
  2025     have cond_uu: "continuous_on (U \<times> U) (\<lambda>(x,y). d x y)"
  2021       apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f')
  2026       apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f')
  2022       apply (simp add: tendsto_within_open_NO_MATCH open_Times \<open>open U\<close>, clarify)
  2027       apply (simp add: tendsto_within_open_NO_MATCH open_Times \<open>open U\<close>, clarify)
  2028     proof -
  2033     proof -
  2029       have "continuous_on U ((\<lambda>(x,y). d x y) \<circ> (\<lambda>z. (w,z)))"
  2034       have "continuous_on U ((\<lambda>(x,y). d x y) \<circ> (\<lambda>z. (w,z)))"
  2030         by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+
  2035         by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+
  2031       then have *: "continuous_on U (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z))"
  2036       then have *: "continuous_on U (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z))"
  2032         by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps)
  2037         by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps)
  2033       have **: "\<And>x. \<lbrakk>x \<in> U; x \<noteq> w\<rbrakk> \<Longrightarrow> (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x"
  2038       have **: "(\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x"
  2034         apply (rule_tac f = "\<lambda>x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within)
  2039         if "x \<in> U" "x \<noteq> w" for x
  2035         apply (rule \<open>open U\<close> derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp: dist_commute)+
  2040       proof (rule_tac f = "\<lambda>x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within)
  2036         done
  2041         show "(\<lambda>x. (f w - f x) / (w - x)) field_differentiable at x"
       
  2042           using that \<open>open U\<close>
       
  2043           by (intro derivative_intros holomorphic_on_imp_differentiable_at [OF holf]; force)
       
  2044       qed (use that \<open>open U\<close> in \<open>auto simp: dist_commute\<close>)
  2037       show ?thesis
  2045       show ?thesis
  2038         unfolding d_def
  2046         unfolding d_def
  2039         apply (rule no_isolated_singularity [OF * _ \<open>open U\<close>, where K = "{w}"])
  2047       proof (rule no_isolated_singularity [OF * _ \<open>open U\<close>])
  2040         apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff \<open>open U\<close> **)
  2048         show "(\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z)) holomorphic_on U - {w}"
  2041         done
  2049           by (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff \<open>open U\<close> **)
       
  2050       qed auto
  2042     qed
  2051     qed
  2043     { fix a b
  2052     { fix a b
  2044       assume abu: "closed_segment a b \<subseteq> U"
  2053       assume abu: "closed_segment a b \<subseteq> U"
  2045       then have "\<And>w. w \<in> U \<Longrightarrow> (\<lambda>z. d z w) contour_integrable_on (linepath a b)"
  2054       have cont_cint_d: "continuous_on U (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
  2046         by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on)
  2055       proof (rule contour_integral_continuous_on_linepath_2D [OF \<open>open U\<close> _ _ abu])
  2047       then have cont_cint_d: "continuous_on U (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
  2056         show "\<And>w. w \<in> U \<Longrightarrow> (\<lambda>z. d z w) contour_integrable_on (linepath a b)"
  2048         apply (rule contour_integral_continuous_on_linepath_2D [OF \<open>open U\<close> _ _ abu])
  2057           by (metis abu hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on)
  2049         apply (auto intro: continuous_on_swap_args cond_uu)
  2058         show "continuous_on (U \<times> U) (\<lambda>(x, y). d y x)"
  2050         done
  2059           by (auto intro: continuous_on_swap_args cond_uu)
       
  2060       qed
  2051       have cont_cint_d\<gamma>: "continuous_on {0..1} ((\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) \<circ> \<gamma>)"
  2061       have cont_cint_d\<gamma>: "continuous_on {0..1} ((\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) \<circ> \<gamma>)"
  2052       proof (rule continuous_on_compose)
  2062       proof (rule continuous_on_compose)
  2053         show "continuous_on {0..1} \<gamma>"
  2063         show "continuous_on {0..1} \<gamma>"
  2054           using \<open>path \<gamma>\<close> path_def by blast
  2064           using \<open>path \<gamma>\<close> path_def by blast
  2055         show "continuous_on (\<gamma> ` {0..1}) (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
  2065         show "continuous_on (\<gamma> ` {0..1}) (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
  2056           using pasz unfolding path_image_def
  2066           using pasz unfolding path_image_def
  2057           by (auto intro!: continuous_on_subset [OF cont_cint_d])
  2067           by (auto intro!: continuous_on_subset [OF cont_cint_d])
  2058       qed
  2068       qed
  2059       have cint_cint: "(\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) contour_integrable_on \<gamma>"
  2069       have "continuous_on {0..1} (\<lambda>x. vector_derivative \<gamma> (at x))"
       
  2070         using pf\<gamma>' by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \<gamma>'])
       
  2071       then      have cint_cint: "(\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) contour_integrable_on \<gamma>"
  2060         apply (simp add: contour_integrable_on)
  2072         apply (simp add: contour_integrable_on)
  2061         apply (rule integrable_continuous_real)
  2073         apply (rule integrable_continuous_real)
  2062         apply (rule continuous_on_mult [OF cont_cint_d\<gamma> [unfolded o_def]])
  2074         by (rule continuous_on_mult [OF cont_cint_d\<gamma> [unfolded o_def]])
  2063         using pf\<gamma>'
       
  2064         by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \<gamma>'])
       
  2065       have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\<lambda>z. contour_integral \<gamma> (d z))"
  2075       have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\<lambda>z. contour_integral \<gamma> (d z))"
  2066         using abu  by (force simp: h_def intro: contour_integral_eq)
  2076         using abu  by (force simp: h_def intro: contour_integral_eq)
  2067       also have "\<dots> =  contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
  2077       also have "\<dots> =  contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
  2068         apply (rule contour_integral_swap)
  2078       proof (rule contour_integral_swap)
  2069         apply (rule continuous_on_subset [OF cond_uu])
  2079         show "continuous_on (path_image (linepath a b) \<times> path_image \<gamma>) (\<lambda>(y1, y2). d y1 y2)"
  2070         using abu pasz \<open>valid_path \<gamma>\<close>
  2080           using abu pasz by (auto intro: continuous_on_subset [OF cond_uu])
  2071         apply (auto intro!: continuous_intros)
  2081         show "continuous_on {0..1} (\<lambda>t. vector_derivative (linepath a b) (at t))"
  2072         by (metis \<gamma>' continuous_on_eq path_def path_polynomial_function pf\<gamma>' vector_derivative_at)
  2082           by (auto intro!: continuous_intros)
       
  2083         show "continuous_on {0..1} (\<lambda>t. vector_derivative \<gamma> (at t))"
       
  2084           by (metis \<gamma>' continuous_on_eq path_def path_polynomial_function pf\<gamma>' vector_derivative_at)
       
  2085       qed (use \<open>valid_path \<gamma>\<close> in auto)
  2073       finally have cint_h_eq:
  2086       finally have cint_h_eq:
  2074           "contour_integral (linepath a b) h =
  2087           "contour_integral (linepath a b) h =
  2075                     contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))" .
  2088                     contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))" .
  2076       note cint_cint cint_h_eq
  2089       note cint_cint cint_h_eq
  2077     } note cint_h = this
  2090     } note cint_h = this
  2089         assume "0 < ee"
  2102         assume "0 < ee"
  2090         show "\<forall>\<^sub>F n in sequentially. \<forall>\<xi>\<in>path_image \<gamma>. cmod (d (a n) \<xi> - d x \<xi>) < ee"
  2103         show "\<forall>\<^sub>F n in sequentially. \<forall>\<xi>\<in>path_image \<gamma>. cmod (d (a n) \<xi> - d x \<xi>) < ee"
  2091         proof -
  2104         proof -
  2092           let ?ddpa = "{(w,z) |w z. w \<in> cball x dd \<and> z \<in> path_image \<gamma>}"
  2105           let ?ddpa = "{(w,z) |w z. w \<in> cball x dd \<and> z \<in> path_image \<gamma>}"
  2093           have "uniformly_continuous_on ?ddpa (\<lambda>(x,y). d x y)"
  2106           have "uniformly_continuous_on ?ddpa (\<lambda>(x,y). d x y)"
  2094             apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]])
  2107           proof (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]])
  2095             using dd pasz \<open>valid_path \<gamma>\<close>
  2108             show "compact {(w, z) |w z. w \<in> cball x dd \<and> z \<in> path_image \<gamma>}"
  2096              apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball)
  2109               using \<open>valid_path \<gamma>\<close>
  2097             done
  2110               by (auto simp: compact_Times compact_valid_path_image simp del: mem_cball)
       
  2111           qed (use dd pasz in auto)
  2098           then obtain kk where "kk>0"
  2112           then obtain kk where "kk>0"
  2099             and kk: "\<And>x x'. \<lbrakk>x \<in> ?ddpa; x' \<in> ?ddpa; dist x' x < kk\<rbrakk> \<Longrightarrow>
  2113             and kk: "\<And>x x'. \<lbrakk>x \<in> ?ddpa; x' \<in> ?ddpa; dist x' x < kk\<rbrakk> \<Longrightarrow>
  2100                              dist ((\<lambda>(x,y). d x y) x') ((\<lambda>(x,y). d x y) x) < ee"
  2114                              dist ((\<lambda>(x,y). d x y) x') ((\<lambda>(x,y). d x y) x) < ee"
  2101             by (rule uniformly_continuous_onE [where e = ee]) (use \<open>0 < ee\<close> in auto)
  2115             by (rule uniformly_continuous_onE [where e = ee]) (use \<open>0 < ee\<close> in auto)
  2102           have kk: "\<lbrakk>norm (w - x) \<le> dd; z \<in> path_image \<gamma>; norm ((w, z) - (x, z)) < kk\<rbrakk> \<Longrightarrow> norm (d w z - d x z) < ee"
  2116           have kk: "\<lbrakk>norm (w - x) \<le> dd; z \<in> path_image \<gamma>; norm ((w, z) - (x, z)) < kk\<rbrakk> \<Longrightarrow> norm (d w z - d x z) < ee"
  2103             for  w z
  2117             for  w z
  2104             using \<open>dd>0\<close> kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm)
  2118             using \<open>dd>0\<close> kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm)
  2105           show ?thesis
  2119           obtain no where "\<forall>n\<ge>no. dist (a n) x < min dd kk"
  2106             using ax unfolding lim_sequentially eventually_sequentially
  2120             using ax unfolding lim_sequentially
  2107             apply (drule_tac x="min dd kk" in spec)
  2121             by (meson \<open>0 < dd\<close> \<open>0 < kk\<close> min_less_iff_conj)
  2108             using \<open>dd > 0\<close> \<open>kk > 0\<close>
  2122           then show ?thesis
  2109             apply (fastforce simp: kk dist_norm)
  2123             using \<open>dd > 0\<close> \<open>kk > 0\<close> by (fastforce simp: eventually_sequentially kk dist_norm)
  2110             done
       
  2111         qed
  2124         qed
  2112       qed
  2125       qed
  2113       have "(\<lambda>n. contour_integral \<gamma> (d (a n))) \<longlonglongrightarrow> contour_integral \<gamma> (d x)"
  2126       have "(\<lambda>n. contour_integral \<gamma> (d (a n))) \<longlonglongrightarrow> contour_integral \<gamma> (d x)"
  2114         by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \<open>valid_path \<gamma>\<close>)
  2127         by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \<open>valid_path \<gamma>\<close>)
  2115       then have tendsto_hx: "(\<lambda>n. contour_integral \<gamma> (d (a n))) \<longlonglongrightarrow> h x"
  2128       then have tendsto_hx: "(\<lambda>n. contour_integral \<gamma> (d (a n))) \<longlonglongrightarrow> h x"
  2136             using  hol_dw holomorphic_on_imp_continuous_on \<open>open U\<close>
  2149             using  hol_dw holomorphic_on_imp_continuous_on \<open>open U\<close>
  2137             by (auto intro!: contour_integrable_holomorphic_simple)
  2150             by (auto intro!: contour_integrable_holomorphic_simple)
  2138           have abc: "closed_segment a b \<subseteq> U"  "closed_segment b c \<subseteq> U"  "closed_segment c a \<subseteq> U"
  2151           have abc: "closed_segment a b \<subseteq> U"  "closed_segment b c \<subseteq> U"  "closed_segment c a \<subseteq> U"
  2139             using that e segments_subset_convex_hull by fastforce+
  2152             using that e segments_subset_convex_hull by fastforce+
  2140           have eq0: "\<And>w. w \<in> U \<Longrightarrow> contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\<lambda>z. d z w) = 0"
  2153           have eq0: "\<And>w. w \<in> U \<Longrightarrow> contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\<lambda>z. d z w) = 0"
  2141             apply (rule contour_integral_unique [OF Cauchy_theorem_triangle])
  2154           proof (rule contour_integral_unique [OF Cauchy_theorem_triangle])
  2142             apply (rule holomorphic_on_subset [OF hol_dw])
  2155             show "\<And>w. w \<in> U \<Longrightarrow> (\<lambda>z. d z w) holomorphic_on convex hull {a, b, c}"
  2143             using e abc_subset by auto
  2156               using e abc_subset by (auto intro: holomorphic_on_subset [OF hol_dw])
       
  2157           qed
  2144           have "contour_integral \<gamma>
  2158           have "contour_integral \<gamma>
  2145                    (\<lambda>x. contour_integral (linepath a b) (\<lambda>z. d z x) +
  2159                    (\<lambda>x. contour_integral (linepath a b) (\<lambda>z. d z x) +
  2146                         (contour_integral (linepath b c) (\<lambda>z. d z x) +
  2160                         (contour_integral (linepath b c) (\<lambda>z. d z x) +
  2147                          contour_integral (linepath c a) (\<lambda>z. d z x)))  =  0"
  2161                          contour_integral (linepath c a) (\<lambda>z. d z x)))  =  0"
  2148             apply (rule contour_integral_eq_0)
  2162             apply (rule contour_integral_eq_0)
  2314         and "0 < r" and "0 < n"
  2328         and "0 < r" and "0 < n"
  2315       shows "norm ((deriv ^^ n) f z) \<le> (fact n) * B0 / r^n"
  2329       shows "norm ((deriv ^^ n) f z) \<le> (fact n) * B0 / r^n"
  2316 proof -
  2330 proof -
  2317   have "0 < B0" using \<open>0 < r\<close> fin [of z]
  2331   have "0 < B0" using \<open>0 < r\<close> fin [of z]
  2318     by (metis ball_eq_empty ex_in_conv fin not_less)
  2332     by (metis ball_eq_empty ex_in_conv fin not_less)
  2319   have le_B0: "\<And>w. cmod (w - z) \<le> r \<Longrightarrow> cmod (f w - y) \<le> B0"
  2333   have le_B0: "cmod (f w - y) \<le> B0" if "cmod (w - z) \<le> r" for w
  2320     apply (rule continuous_on_closure_norm_le [of "ball z r" "\<lambda>w. f w - y"])
  2334   proof (rule continuous_on_closure_norm_le [of "ball z r" "\<lambda>w. f w - y"], use \<open>0 < r\<close> in simp_all)
  2321     apply (auto simp: \<open>0 < r\<close>  dist_norm norm_minus_commute)
  2335     show "continuous_on (cball z r) (\<lambda>w. f w - y)"
  2322     apply (rule continuous_intros contf)+
  2336       by (intro continuous_intros contf)
  2323     using fin apply (simp add: dist_commute dist_norm less_eq_real_def)
  2337     show "dist z w \<le> r"
  2324     done
  2338       by (simp add: dist_commute dist_norm that)
       
  2339     qed (use fin in \<open>auto simp: dist_norm less_eq_real_def norm_minus_commute\<close>)
  2325   have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w) z - (deriv ^^ n) (\<lambda>w. y) z"
  2340   have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w) z - (deriv ^^ n) (\<lambda>w. y) z"
  2326     using \<open>0 < n\<close> by simp
  2341     using \<open>0 < n\<close> by simp
  2327   also have "... = (deriv ^^ n) (\<lambda>w. f w - y) z"
  2342   also have "... = (deriv ^^ n) (\<lambda>w. f w - y) z"
  2328     by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \<open>0 < r\<close>)
  2343     by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \<open>0 < r\<close>)
  2329   finally have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w - y) z" .
  2344   finally have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w - y) z" .
  2360   obtain x where "norm (\<xi>-x) = r"
  2375   obtain x where "norm (\<xi>-x) = r"
  2361     by (metis abs_of_nonneg add_diff_cancel_left' \<open>0 < r\<close> diff_add_cancel
  2376     by (metis abs_of_nonneg add_diff_cancel_left' \<open>0 < r\<close> diff_add_cancel
  2362                  dual_order.strict_implies_order norm_of_real)
  2377                  dual_order.strict_implies_order norm_of_real)
  2363   then have "0 \<le> B"
  2378   then have "0 \<le> B"
  2364     by (metis nof norm_not_less_zero not_le order_trans)
  2379     by (metis nof norm_not_less_zero not_le order_trans)
  2365   have  "((\<lambda>u. f u / (u - \<xi>) ^ Suc n) has_contour_integral (2 * pi) * \<i> / fact n * (deriv ^^ n) f \<xi>)
  2380   have "\<xi> \<in> ball \<xi> r"
       
  2381     using \<open>0 < r\<close> by simp
       
  2382   then have  "((\<lambda>u. f u / (u-\<xi>) ^ Suc n) has_contour_integral (2 * pi) * \<i> / fact n * (deriv ^^ n) f \<xi>)
  2366          (circlepath \<xi> r)"
  2383          (circlepath \<xi> r)"
  2367     apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf])
  2384     by (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf])
  2368     using \<open>0 < r\<close> by simp
  2385   have "norm ((2 * pi * \<i>)/(fact n) * (deriv ^^ n) f \<xi>) \<le> (B / r^(Suc n)) * (2 * pi * r)"
  2369   then have "norm ((2 * pi * \<i>)/(fact n) * (deriv ^^ n) f \<xi>) \<le> (B / r^(Suc n)) * (2 * pi * r)"
  2386   proof (rule has_contour_integral_bound_circlepath)
  2370     apply (rule has_contour_integral_bound_circlepath)
  2387     have "\<xi> \<in> ball \<xi> r"
  2371     using \<open>0 \<le> B\<close> \<open>0 < r\<close>
  2388       using \<open>0 < r\<close> by simp
  2372     apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc)
  2389     then show  "((\<lambda>u. f u / (u-\<xi>) ^ Suc n) has_contour_integral (2 * pi) * \<i> / fact n * (deriv ^^ n) f \<xi>)
  2373     done
  2390          (circlepath \<xi> r)"
       
  2391       by (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf])
       
  2392     show "\<And>x. cmod (x-\<xi>) = r \<Longrightarrow> cmod (f x / (x-\<xi>) ^ Suc n) \<le> B / r ^ Suc n"
       
  2393       using \<open>0 \<le> B\<close> \<open>0 < r\<close>
       
  2394       by (simp add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc)
       
  2395   qed (use \<open>0 \<le> B\<close> \<open>0 < r\<close> in auto)
  2374   then show ?thesis using \<open>0 < r\<close>
  2396   then show ?thesis using \<open>0 < r\<close>
  2375     by (simp add: norm_divide norm_mult field_simps)
  2397     by (simp add: norm_divide norm_mult field_simps)
  2376 qed
  2398 qed
  2377 
  2399 
  2378 lemma Liouville_polynomial:
  2400 lemma Liouville_polynomial:
  2390     by (simp add: eventually_at_infinity f0) meson
  2412     by (simp add: eventually_at_infinity f0) meson
  2391   show ?thesis by simp
  2413   show ?thesis by simp
  2392 next
  2414 next
  2393   assume "0 < B"
  2415   assume "0 < B"
  2394   have "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * (\<xi> - 0)^k) sums f \<xi>)"
  2416   have "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * (\<xi> - 0)^k) sums f \<xi>)"
  2395     apply (rule holomorphic_power_series [where r = "norm \<xi> + 1"])
  2417   proof (rule holomorphic_power_series [where r = "norm \<xi> + 1"])
  2396     using holf holomorphic_on_subset apply auto
  2418     show "f holomorphic_on ball 0 (cmod \<xi> + 1)" "\<xi> \<in> ball 0 (cmod \<xi> + 1)"
  2397     done
  2419       using holf holomorphic_on_subset by auto
       
  2420   qed
  2398   then have sumsf: "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * \<xi>^k) sums f \<xi>)" by simp
  2421   then have sumsf: "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * \<xi>^k) sums f \<xi>)" by simp
  2399   have "(deriv ^^ k) f 0 / fact k * \<xi> ^ k = 0" if "k>n" for k
  2422   have "(deriv ^^ k) f 0 / fact k * \<xi> ^ k = 0" if "k>n" for k
  2400   proof (cases "(deriv ^^ k) f 0 = 0")
  2423   proof (cases "(deriv ^^ k) f 0 = 0")
  2401     case True then show ?thesis by simp
  2424     case True then show ?thesis by simp
  2402   next
  2425   next
  2418     then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w"
  2441     then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w"
  2419       by (metis norm_of_real w_def)
  2442       by (metis norm_of_real w_def)
  2420     then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)"
  2443     then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)"
  2421       using False by (simp add: field_split_simps mult.commute split: if_split_asm)
  2444       using False by (simp add: field_split_simps mult.commute split: if_split_asm)
  2422     also have "... \<le> fact k * (B * norm w ^ n) / norm w ^ k"
  2445     also have "... \<le> fact k * (B * norm w ^ n) / norm w ^ k"
  2423       apply (rule Cauchy_inequality)
  2446     proof (rule Cauchy_inequality)
  2424          using holf holomorphic_on_subset apply force
  2447       show "f holomorphic_on ball 0 (cmod w)"
  2425         using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast
  2448         using holf holomorphic_on_subset by force
  2426        using \<open>w \<noteq> 0\<close> apply simp
  2449       show "continuous_on (cball 0 (cmod w)) f"
  2427        by (metis nof wgeA dist_0_norm dist_norm)
  2450         using holf holomorphic_on_imp_continuous_on holomorphic_on_subset by blast
       
  2451       show "\<And>x. cmod (0 - x) = cmod w \<Longrightarrow> cmod (f x) \<le> B * cmod w ^ n"
       
  2452         by (metis nof wgeA dist_0_norm dist_norm)
       
  2453     qed (use \<open>w \<noteq> 0\<close> in auto)
  2428     also have "... = fact k * (B * 1 / cmod w ^ (k-n))"
  2454     also have "... = fact k * (B * 1 / cmod w ^ (k-n))"
  2429       apply (simp only: mult_cancel_left times_divide_eq_right [symmetric])
  2455     proof -
  2430       using \<open>k>n\<close> \<open>w \<noteq> 0\<close> \<open>0 < B\<close> apply (simp add: field_split_simps semiring_normalization_rules)
  2456       have "cmod w ^ n / cmod w ^ k = 1 / cmod w ^ (k - n)"
  2431       done
  2457         using \<open>k>n\<close> \<open>w \<noteq> 0\<close> \<open>0 < B\<close> by (simp add: field_split_simps semiring_normalization_rules)
       
  2458       then show ?thesis
       
  2459         by (metis times_divide_eq_right)
       
  2460     qed
  2432     also have "... = fact k * B / cmod w ^ (k-n)"
  2461     also have "... = fact k * B / cmod w ^ (k-n)"
  2433       by simp
  2462       by simp
  2434     finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" .
  2463     finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" .
  2435     then have "1 / cmod w < 1 / cmod w ^ (k - n)"
  2464     then have "1 / cmod w < 1 / cmod w ^ (k - n)"
  2436       by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos)
  2465       by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos)
  2464 text\<open>A holomorphic function f has only isolated zeros unless f is 0.\<close>
  2493 text\<open>A holomorphic function f has only isolated zeros unless f is 0.\<close>
  2465 
  2494 
  2466 lemma powser_0_nonzero:
  2495 lemma powser_0_nonzero:
  2467   fixes a :: "nat \<Rightarrow> 'a::{real_normed_field,banach}"
  2496   fixes a :: "nat \<Rightarrow> 'a::{real_normed_field,banach}"
  2468   assumes r: "0 < r"
  2497   assumes r: "0 < r"
  2469       and sm: "\<And>x. norm (x - \<xi>) < r \<Longrightarrow> (\<lambda>n. a n * (x - \<xi>) ^ n) sums (f x)"
  2498       and sm: "\<And>x. norm (x-\<xi>) < r \<Longrightarrow> (\<lambda>n. a n * (x-\<xi>) ^ n) sums (f x)"
  2470       and [simp]: "f \<xi> = 0"
  2499       and [simp]: "f \<xi> = 0"
  2471       and m0: "a m \<noteq> 0" and "m>0"
  2500       and m0: "a m \<noteq> 0" and "m>0"
  2472   obtains s where "0 < s" and "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0"
  2501   obtains s where "0 < s" and "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0"
  2473 proof -
  2502 proof -
  2474   have "r \<le> conv_radius a"
  2503   have "r \<le> conv_radius a"
  2475     using sm sums_summable by (auto simp: le_conv_radius_iff [where \<xi>=\<xi>])
  2504     using sm sums_summable by (auto simp: le_conv_radius_iff [where \<xi>=\<xi>])
  2476   obtain m where am: "a m \<noteq> 0" and az [simp]: "(\<And>n. n<m \<Longrightarrow> a n = 0)"
  2505   obtain m where am: "a m \<noteq> 0" and az [simp]: "(\<And>n. n<m \<Longrightarrow> a n = 0)"
  2477     apply (rule_tac m = "LEAST n. a n \<noteq> 0" in that)
  2506   proof
  2478     using m0
  2507     show "a (LEAST n. a n \<noteq> 0) \<noteq> 0"
  2479     apply (rule LeastI2)
  2508       by (metis (mono_tags, lifting) m0 LeastI)
  2480     apply (fastforce intro:  dest!: not_less_Least)+
  2509   qed (fastforce dest!: not_less_Least)
  2481     done
       
  2482   define b where "b i = a (i+m) / a m" for i
  2510   define b where "b i = a (i+m) / a m" for i
  2483   define g where "g x = suminf (\<lambda>i. b i * (x - \<xi>) ^ i)" for x
  2511   define g where "g x = suminf (\<lambda>i. b i * (x-\<xi>) ^ i)" for x
  2484   have [simp]: "b 0 = 1"
  2512   have [simp]: "b 0 = 1"
  2485     by (simp add: am b_def)
  2513     by (simp add: am b_def)
  2486   { fix x::'a
  2514   { fix x::'a
  2487     assume "norm (x - \<xi>) < r"
  2515     assume "norm (x-\<xi>) < r"
  2488     then have "(\<lambda>n. (a m * (x - \<xi>)^m) * (b n * (x - \<xi>)^n)) sums (f x)"
  2516     then have "(\<lambda>n. (a m * (x-\<xi>)^m) * (b n * (x-\<xi>)^n)) sums (f x)"
  2489       using am az sm sums_zero_iff_shift [of m "(\<lambda>n. a n * (x - \<xi>) ^ n)" "f x"]
  2517       using am az sm sums_zero_iff_shift [of m "(\<lambda>n. a n * (x-\<xi>) ^ n)" "f x"]
  2490       by (simp add: b_def monoid_mult_class.power_add algebra_simps)
  2518       by (simp add: b_def monoid_mult_class.power_add algebra_simps)
  2491     then have "x \<noteq> \<xi> \<Longrightarrow> (\<lambda>n. b n * (x - \<xi>)^n) sums (f x / (a m * (x - \<xi>)^m))"
  2519     then have "x \<noteq> \<xi> \<Longrightarrow> (\<lambda>n. b n * (x-\<xi>)^n) sums (f x / (a m * (x-\<xi>)^m))"
  2492       using am by (simp add: sums_mult_D)
  2520       using am by (simp add: sums_mult_D)
  2493   } note bsums = this
  2521   } note bsums = this
  2494   then have  "norm (x - \<xi>) < r \<Longrightarrow> summable (\<lambda>n. b n * (x - \<xi>)^n)" for x
  2522   then have  "norm (x-\<xi>) < r \<Longrightarrow> summable (\<lambda>n. b n * (x-\<xi>)^n)" for x
  2495     using sums_summable by (cases "x=\<xi>") auto
  2523     using sums_summable by (cases "x=\<xi>") auto
  2496   then have "r \<le> conv_radius b"
  2524   then have "r \<le> conv_radius b"
  2497     by (simp add: le_conv_radius_iff [where \<xi>=\<xi>])
  2525     by (simp add: le_conv_radius_iff [where \<xi>=\<xi>])
  2498   then have "r/2 < conv_radius b"
  2526   then have "r/2 < conv_radius b"
  2499     using not_le order_trans r by fastforce
  2527     using not_le order_trans r by fastforce
  2500   then have "continuous_on (cball \<xi> (r/2)) g"
  2528   then have "continuous_on (cball \<xi> (r/2)) g"
  2501     using powser_continuous_suminf [of "r/2" b \<xi>] by (simp add: g_def)
  2529     using powser_continuous_suminf [of "r/2" b \<xi>] by (simp add: g_def)
  2502   then obtain s where "s>0"  "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> dist (g x) (g \<xi>) < 1/2"
  2530   then obtain s where "s>0"  "\<And>x. \<lbrakk>norm (x-\<xi>) \<le> s; norm (x-\<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> dist (g x) (g \<xi>) < 1/2"
  2503     apply (rule continuous_onE [where x=\<xi> and e = "1/2"])
  2531   proof (rule continuous_onE)
  2504     using r apply (auto simp: norm_minus_commute dist_norm)
  2532     show "\<xi> \<in> cball \<xi> (r / 2)" "1/2 > (0::real)"
  2505     done
  2533       using r by auto
       
  2534   qed (auto simp: dist_commute dist_norm)
  2506   moreover have "g \<xi> = 1"
  2535   moreover have "g \<xi> = 1"
  2507     by (simp add: g_def)
  2536     by (simp add: g_def)
  2508   ultimately have gnz: "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> (g x) \<noteq> 0"
  2537   ultimately have gnz: "\<And>x. \<lbrakk>norm (x-\<xi>) \<le> s; norm (x-\<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> (g x) \<noteq> 0"
  2509     by fastforce
  2538     by fastforce
  2510   have "f x \<noteq> 0" if "x \<noteq> \<xi>" "norm (x - \<xi>) \<le> s" "norm (x - \<xi>) \<le> r/2" for x
  2539   have "f x \<noteq> 0" if "x \<noteq> \<xi>" "norm (x-\<xi>) \<le> s" "norm (x-\<xi>) \<le> r/2" for x
  2511     using bsums [of x] that gnz [of x]
  2540     using bsums [of x] that gnz [of x] r sums_iff unfolding g_def by fastforce
  2512     apply (auto simp: g_def)
       
  2513     using r sums_iff by fastforce
       
  2514   then show ?thesis
  2541   then show ?thesis
  2515     apply (rule_tac s="min s (r/2)" in that)
  2542     apply (rule_tac s="min s (r/2)" in that)
  2516     using \<open>0 < r\<close> \<open>0 < s\<close> by (auto simp: dist_commute dist_norm)
  2543     using \<open>0 < r\<close> \<open>0 < s\<close> by (auto simp: dist_commute dist_norm)
  2517 qed
  2544 qed
  2518 
  2545 
  2778   case False
  2805   case False
  2779   show ?thesis unfolding fps_tan_def
  2806   show ?thesis unfolding fps_tan_def
  2780   proof (subst eval_fps_divide'[where r = "pi / (2 * norm c)"])
  2807   proof (subst eval_fps_divide'[where r = "pi / (2 * norm c)"])
  2781     fix z :: complex assume "z \<in> eball 0 (pi / (2 * norm c))"
  2808     fix z :: complex assume "z \<in> eball 0 (pi / (2 * norm c))"
  2782     with cos_eq_zero_imp_norm_ge[of "c*z"] assms 
  2809     with cos_eq_zero_imp_norm_ge[of "c*z"] assms 
  2783       show "eval_fps (fps_cos  c) z \<noteq> 0" using False by (auto simp: norm_mult field_simps)
  2810     show "eval_fps (fps_cos  c) z \<noteq> 0" using False by (auto simp: norm_mult field_simps)
  2784     qed (insert False assms, auto simp: field_simps tan_def)
  2811   qed (insert False assms, auto simp: field_simps tan_def)
  2785   qed simp_all
  2812 qed simp_all
  2786 
  2813 
  2787 end
  2814 end