src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy
changeset 63594 bd218a9320b5
parent 63593 bbcb05504fdc
child 63595 aca2659ebba7
equal deleted inserted replaced
63593:bbcb05504fdc 63594:bd218a9320b5
     1 section \<open>Complex path integrals and Cauchy's integral theorem\<close>
       
     2 
       
     3 text\<open>By John Harrison et al.  Ported from HOL Light by L C Paulson (2015)\<close>
       
     4 
       
     5 theory Cauchy_Integral_Thm
       
     6 imports Complex_Transcendental Weierstrass Ordered_Euclidean_Space
       
     7 begin
       
     8 
       
     9 subsection\<open>Homeomorphisms of arc images\<close>
       
    10 
       
    11 lemma homeomorphism_arc:
       
    12   fixes g :: "real \<Rightarrow> 'a::t2_space"
       
    13   assumes "arc g"
       
    14   obtains h where "homeomorphism {0..1} (path_image g) g h"
       
    15 using assms by (force simp add: arc_def homeomorphism_compact path_def path_image_def)
       
    16 
       
    17 lemma homeomorphic_arc_image_interval:
       
    18   fixes g :: "real \<Rightarrow> 'a::t2_space" and a::real
       
    19   assumes "arc g" "a < b"
       
    20   shows "(path_image g) homeomorphic {a..b}"
       
    21 proof -
       
    22   have "(path_image g) homeomorphic {0..1::real}"
       
    23     by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc)
       
    24   also have "... homeomorphic {a..b}"
       
    25     using assms by (force intro: homeomorphic_closed_intervals_real)
       
    26   finally show ?thesis .
       
    27 qed
       
    28 
       
    29 lemma homeomorphic_arc_images:
       
    30   fixes g :: "real \<Rightarrow> 'a::t2_space" and h :: "real \<Rightarrow> 'b::t2_space"
       
    31   assumes "arc g" "arc h"
       
    32   shows "(path_image g) homeomorphic (path_image h)"
       
    33 proof -
       
    34   have "(path_image g) homeomorphic {0..1::real}"
       
    35     by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc)
       
    36   also have "... homeomorphic (path_image h)"
       
    37     by (meson assms homeomorphic_def homeomorphism_arc)
       
    38   finally show ?thesis .
       
    39 qed
       
    40 
       
    41 subsection \<open>Piecewise differentiable functions\<close>
       
    42 
       
    43 definition piecewise_differentiable_on
       
    44            (infixr "piecewise'_differentiable'_on" 50)
       
    45   where "f piecewise_differentiable_on i  \<equiv>
       
    46            continuous_on i f \<and>
       
    47            (\<exists>s. finite s \<and> (\<forall>x \<in> i - s. f differentiable (at x within i)))"
       
    48 
       
    49 lemma piecewise_differentiable_on_imp_continuous_on:
       
    50     "f piecewise_differentiable_on s \<Longrightarrow> continuous_on s f"
       
    51 by (simp add: piecewise_differentiable_on_def)
       
    52 
       
    53 lemma piecewise_differentiable_on_subset:
       
    54     "f piecewise_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_differentiable_on t"
       
    55   using continuous_on_subset
       
    56   unfolding piecewise_differentiable_on_def
       
    57   apply safe
       
    58   apply (blast intro: elim: continuous_on_subset)
       
    59   by (meson Diff_iff differentiable_within_subset subsetCE)
       
    60 
       
    61 lemma differentiable_on_imp_piecewise_differentiable:
       
    62   fixes a:: "'a::{linorder_topology,real_normed_vector}"
       
    63   shows "f differentiable_on {a..b} \<Longrightarrow> f piecewise_differentiable_on {a..b}"
       
    64   apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on)
       
    65   apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def)
       
    66   done
       
    67 
       
    68 lemma differentiable_imp_piecewise_differentiable:
       
    69     "(\<And>x. x \<in> s \<Longrightarrow> f differentiable (at x within s))
       
    70          \<Longrightarrow> f piecewise_differentiable_on s"
       
    71 by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def
       
    72          intro: differentiable_within_subset)
       
    73 
       
    74 lemma piecewise_differentiable_const [iff]: "(\<lambda>x. z) piecewise_differentiable_on s"
       
    75   by (simp add: differentiable_imp_piecewise_differentiable)
       
    76 
       
    77 lemma piecewise_differentiable_compose:
       
    78     "\<lbrakk>f piecewise_differentiable_on s; g piecewise_differentiable_on (f ` s);
       
    79       \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
       
    80       \<Longrightarrow> (g o f) piecewise_differentiable_on s"
       
    81   apply (simp add: piecewise_differentiable_on_def, safe)
       
    82   apply (blast intro: continuous_on_compose2)
       
    83   apply (rename_tac A B)
       
    84   apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI)
       
    85   apply (blast intro: differentiable_chain_within)
       
    86   done
       
    87 
       
    88 lemma piecewise_differentiable_affine:
       
    89   fixes m::real
       
    90   assumes "f piecewise_differentiable_on ((\<lambda>x. m *\<^sub>R x + c) ` s)"
       
    91   shows "(f o (\<lambda>x. m *\<^sub>R x + c)) piecewise_differentiable_on s"
       
    92 proof (cases "m = 0")
       
    93   case True
       
    94   then show ?thesis
       
    95     unfolding o_def
       
    96     by (force intro: differentiable_imp_piecewise_differentiable differentiable_const)
       
    97 next
       
    98   case False
       
    99   show ?thesis
       
   100     apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable])
       
   101     apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+
       
   102     done
       
   103 qed
       
   104 
       
   105 lemma piecewise_differentiable_cases:
       
   106   fixes c::real
       
   107   assumes "f piecewise_differentiable_on {a..c}"
       
   108           "g piecewise_differentiable_on {c..b}"
       
   109            "a \<le> c" "c \<le> b" "f c = g c"
       
   110   shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_differentiable_on {a..b}"
       
   111 proof -
       
   112   obtain s t where st: "finite s" "finite t"
       
   113                        "\<forall>x\<in>{a..c} - s. f differentiable at x within {a..c}"
       
   114                        "\<forall>x\<in>{c..b} - t. g differentiable at x within {c..b}"
       
   115     using assms
       
   116     by (auto simp: piecewise_differentiable_on_def)
       
   117   have finabc: "finite ({a,b,c} \<union> (s \<union> t))"
       
   118     by (metis \<open>finite s\<close> \<open>finite t\<close> finite_Un finite_insert finite.emptyI)
       
   119   have "continuous_on {a..c} f" "continuous_on {c..b} g"
       
   120     using assms piecewise_differentiable_on_def by auto
       
   121   then have "continuous_on {a..b} (\<lambda>x. if x \<le> c then f x else g x)"
       
   122     using continuous_on_cases [OF closed_real_atLeastAtMost [of a c],
       
   123                                OF closed_real_atLeastAtMost [of c b],
       
   124                                of f g "\<lambda>x. x\<le>c"]  assms
       
   125     by (force simp: ivl_disj_un_two_touch)
       
   126   moreover
       
   127   { fix x
       
   128     assume x: "x \<in> {a..b} - ({a,b,c} \<union> (s \<union> t))"
       
   129     have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg")
       
   130     proof (cases x c rule: le_cases)
       
   131       case le show ?diff_fg
       
   132         apply (rule differentiable_transform_within [where d = "dist x c" and f = f])
       
   133         using x le st
       
   134         apply (simp_all add: dist_real_def)
       
   135         apply (rule differentiable_at_withinI)
       
   136         apply (rule differentiable_within_open [where s = "{a<..<c} - s", THEN iffD1], simp_all)
       
   137         apply (blast intro: open_greaterThanLessThan finite_imp_closed)
       
   138         apply (force elim!: differentiable_subset)+
       
   139         done
       
   140     next
       
   141       case ge show ?diff_fg
       
   142         apply (rule differentiable_transform_within [where d = "dist x c" and f = g])
       
   143         using x ge st
       
   144         apply (simp_all add: dist_real_def)
       
   145         apply (rule differentiable_at_withinI)
       
   146         apply (rule differentiable_within_open [where s = "{c<..<b} - t", THEN iffD1], simp_all)
       
   147         apply (blast intro: open_greaterThanLessThan finite_imp_closed)
       
   148         apply (force elim!: differentiable_subset)+
       
   149         done
       
   150     qed
       
   151   }
       
   152   then have "\<exists>s. finite s \<and>
       
   153                  (\<forall>x\<in>{a..b} - s. (\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b})"
       
   154     by (meson finabc)
       
   155   ultimately show ?thesis
       
   156     by (simp add: piecewise_differentiable_on_def)
       
   157 qed
       
   158 
       
   159 lemma piecewise_differentiable_neg:
       
   160     "f piecewise_differentiable_on s \<Longrightarrow> (\<lambda>x. -(f x)) piecewise_differentiable_on s"
       
   161   by (auto simp: piecewise_differentiable_on_def continuous_on_minus)
       
   162 
       
   163 lemma piecewise_differentiable_add:
       
   164   assumes "f piecewise_differentiable_on i"
       
   165           "g piecewise_differentiable_on i"
       
   166     shows "(\<lambda>x. f x + g x) piecewise_differentiable_on i"
       
   167 proof -
       
   168   obtain s t where st: "finite s" "finite t"
       
   169                        "\<forall>x\<in>i - s. f differentiable at x within i"
       
   170                        "\<forall>x\<in>i - t. g differentiable at x within i"
       
   171     using assms by (auto simp: piecewise_differentiable_on_def)
       
   172   then have "finite (s \<union> t) \<and> (\<forall>x\<in>i - (s \<union> t). (\<lambda>x. f x + g x) differentiable at x within i)"
       
   173     by auto
       
   174   moreover have "continuous_on i f" "continuous_on i g"
       
   175     using assms piecewise_differentiable_on_def by auto
       
   176   ultimately show ?thesis
       
   177     by (auto simp: piecewise_differentiable_on_def continuous_on_add)
       
   178 qed
       
   179 
       
   180 lemma piecewise_differentiable_diff:
       
   181     "\<lbrakk>f piecewise_differentiable_on s;  g piecewise_differentiable_on s\<rbrakk>
       
   182      \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_differentiable_on s"
       
   183   unfolding diff_conv_add_uminus
       
   184   by (metis piecewise_differentiable_add piecewise_differentiable_neg)
       
   185 
       
   186 lemma continuous_on_joinpaths_D1:
       
   187     "continuous_on {0..1} (g1 +++ g2) \<Longrightarrow> continuous_on {0..1} g1"
       
   188   apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (op*(inverse 2))"])
       
   189   apply (rule continuous_intros | simp)+
       
   190   apply (auto elim!: continuous_on_subset simp: joinpaths_def)
       
   191   done
       
   192 
       
   193 lemma continuous_on_joinpaths_D2:
       
   194     "\<lbrakk>continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> continuous_on {0..1} g2"
       
   195   apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (\<lambda>x. inverse 2*x + 1/2)"])
       
   196   apply (rule continuous_intros | simp)+
       
   197   apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def)
       
   198   done
       
   199 
       
   200 lemma piecewise_differentiable_D1:
       
   201     "(g1 +++ g2) piecewise_differentiable_on {0..1} \<Longrightarrow> g1 piecewise_differentiable_on {0..1}"
       
   202   apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D1)
       
   203   apply (rule_tac x="insert 1 ((op*2)`s)" in exI)
       
   204   apply simp
       
   205   apply (intro ballI)
       
   206   apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))"
       
   207        in differentiable_transform_within)
       
   208   apply (auto simp: dist_real_def joinpaths_def)
       
   209   apply (rule differentiable_chain_within derivative_intros | simp)+
       
   210   apply (rule differentiable_subset)
       
   211   apply (force simp:)+
       
   212   done
       
   213 
       
   214 lemma piecewise_differentiable_D2:
       
   215     "\<lbrakk>(g1 +++ g2) piecewise_differentiable_on {0..1}; pathfinish g1 = pathstart g2\<rbrakk>
       
   216     \<Longrightarrow> g2 piecewise_differentiable_on {0..1}"
       
   217   apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D2)
       
   218   apply (rule_tac x="insert 0 ((\<lambda>x. 2*x-1)`s)" in exI)
       
   219   apply simp
       
   220   apply (intro ballI)
       
   221   apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)"
       
   222           in differentiable_transform_within)
       
   223   apply (auto simp: dist_real_def joinpaths_def abs_if field_simps split: if_split_asm)
       
   224   apply (rule differentiable_chain_within derivative_intros | simp)+
       
   225   apply (rule differentiable_subset)
       
   226   apply (force simp: divide_simps)+
       
   227   done
       
   228 
       
   229 
       
   230 subsubsection\<open>The concept of continuously differentiable\<close>
       
   231 
       
   232 text \<open>
       
   233 John Harrison writes as follows:
       
   234 
       
   235 ``The usual assumption in complex analysis texts is that a path \<open>\<gamma>\<close> should be piecewise
       
   236 continuously differentiable, which ensures that the path integral exists at least for any continuous
       
   237 f, since all piecewise continuous functions are integrable. However, our notion of validity is
       
   238 weaker, just piecewise differentiability... [namely] continuity plus differentiability except on a
       
   239 finite set ... [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to
       
   240 the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this
       
   241 can integrate all derivatives.''
       
   242 
       
   243 "Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec.
       
   244 Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165.
       
   245 
       
   246 And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably
       
   247 difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem
       
   248 asserting that all derivatives can be integrated before we can adopt Harrison's choice.\<close>
       
   249 
       
   250 definition C1_differentiable_on :: "(real \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> real set \<Rightarrow> bool"
       
   251            (infix "C1'_differentiable'_on" 50)
       
   252   where
       
   253   "f C1_differentiable_on s \<longleftrightarrow>
       
   254    (\<exists>D. (\<forall>x \<in> s. (f has_vector_derivative (D x)) (at x)) \<and> continuous_on s D)"
       
   255 
       
   256 lemma C1_differentiable_on_eq:
       
   257     "f C1_differentiable_on s \<longleftrightarrow>
       
   258      (\<forall>x \<in> s. f differentiable at x) \<and> continuous_on s (\<lambda>x. vector_derivative f (at x))"
       
   259   unfolding C1_differentiable_on_def
       
   260   apply safe
       
   261   using differentiable_def has_vector_derivative_def apply blast
       
   262   apply (erule continuous_on_eq)
       
   263   using vector_derivative_at apply fastforce
       
   264   using vector_derivative_works apply fastforce
       
   265   done
       
   266 
       
   267 lemma C1_differentiable_on_subset:
       
   268   "f C1_differentiable_on t \<Longrightarrow> s \<subseteq> t \<Longrightarrow> f C1_differentiable_on s"
       
   269   unfolding C1_differentiable_on_def  continuous_on_eq_continuous_within
       
   270   by (blast intro:  continuous_within_subset)
       
   271 
       
   272 lemma C1_differentiable_compose:
       
   273     "\<lbrakk>f C1_differentiable_on s; g C1_differentiable_on (f ` s);
       
   274       \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
       
   275       \<Longrightarrow> (g o f) C1_differentiable_on s"
       
   276   apply (simp add: C1_differentiable_on_eq, safe)
       
   277    using differentiable_chain_at apply blast
       
   278   apply (rule continuous_on_eq [of _ "\<lambda>x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"])
       
   279    apply (rule Limits.continuous_on_scaleR, assumption)
       
   280    apply (metis (mono_tags, lifting) continuous_on_eq continuous_at_imp_continuous_on continuous_on_compose differentiable_imp_continuous_within o_def)
       
   281   by (simp add: vector_derivative_chain_at)
       
   282 
       
   283 lemma C1_diff_imp_diff: "f C1_differentiable_on s \<Longrightarrow> f differentiable_on s"
       
   284   by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on)
       
   285 
       
   286 lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\<lambda>x. x) C1_differentiable_on s"
       
   287   by (auto simp: C1_differentiable_on_eq continuous_on_const)
       
   288 
       
   289 lemma C1_differentiable_on_const [simp, derivative_intros]: "(\<lambda>z. a) C1_differentiable_on s"
       
   290   by (auto simp: C1_differentiable_on_eq continuous_on_const)
       
   291 
       
   292 lemma C1_differentiable_on_add [simp, derivative_intros]:
       
   293   "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x + g x) C1_differentiable_on s"
       
   294   unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
       
   295 
       
   296 lemma C1_differentiable_on_minus [simp, derivative_intros]:
       
   297   "f C1_differentiable_on s \<Longrightarrow> (\<lambda>x. - f x) C1_differentiable_on s"
       
   298   unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
       
   299 
       
   300 lemma C1_differentiable_on_diff [simp, derivative_intros]:
       
   301   "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x - g x) C1_differentiable_on s"
       
   302   unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
       
   303 
       
   304 lemma C1_differentiable_on_mult [simp, derivative_intros]:
       
   305   fixes f g :: "real \<Rightarrow> 'a :: real_normed_algebra"
       
   306   shows "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x * g x) C1_differentiable_on s"
       
   307   unfolding C1_differentiable_on_eq
       
   308   by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within)
       
   309 
       
   310 lemma C1_differentiable_on_scaleR [simp, derivative_intros]:
       
   311   "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x *\<^sub>R g x) C1_differentiable_on s"
       
   312   unfolding C1_differentiable_on_eq
       
   313   by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+
       
   314 
       
   315 
       
   316 definition piecewise_C1_differentiable_on
       
   317            (infixr "piecewise'_C1'_differentiable'_on" 50)
       
   318   where "f piecewise_C1_differentiable_on i  \<equiv>
       
   319            continuous_on i f \<and>
       
   320            (\<exists>s. finite s \<and> (f C1_differentiable_on (i - s)))"
       
   321 
       
   322 lemma C1_differentiable_imp_piecewise:
       
   323     "f C1_differentiable_on s \<Longrightarrow> f piecewise_C1_differentiable_on s"
       
   324   by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within)
       
   325 
       
   326 lemma piecewise_C1_imp_differentiable:
       
   327     "f piecewise_C1_differentiable_on i \<Longrightarrow> f piecewise_differentiable_on i"
       
   328   by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def
       
   329            C1_differentiable_on_def differentiable_def has_vector_derivative_def
       
   330            intro: has_derivative_at_within)
       
   331 
       
   332 lemma piecewise_C1_differentiable_compose:
       
   333     "\<lbrakk>f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on (f ` s);
       
   334       \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
       
   335       \<Longrightarrow> (g o f) piecewise_C1_differentiable_on s"
       
   336   apply (simp add: piecewise_C1_differentiable_on_def, safe)
       
   337   apply (blast intro: continuous_on_compose2)
       
   338   apply (rename_tac A B)
       
   339   apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI)
       
   340   apply (rule conjI, blast)
       
   341   apply (rule C1_differentiable_compose)
       
   342   apply (blast intro: C1_differentiable_on_subset)
       
   343   apply (blast intro: C1_differentiable_on_subset)
       
   344   by (simp add: Diff_Int_distrib2)
       
   345 
       
   346 lemma piecewise_C1_differentiable_on_subset:
       
   347     "f piecewise_C1_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_C1_differentiable_on t"
       
   348   by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset)
       
   349 
       
   350 lemma C1_differentiable_imp_continuous_on:
       
   351   "f C1_differentiable_on s \<Longrightarrow> continuous_on s f"
       
   352   unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within
       
   353   using differentiable_at_withinI differentiable_imp_continuous_within by blast
       
   354 
       
   355 lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}"
       
   356   unfolding C1_differentiable_on_def
       
   357   by auto
       
   358 
       
   359 lemma piecewise_C1_differentiable_affine:
       
   360   fixes m::real
       
   361   assumes "f piecewise_C1_differentiable_on ((\<lambda>x. m * x + c) ` s)"
       
   362   shows "(f o (\<lambda>x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on s"
       
   363 proof (cases "m = 0")
       
   364   case True
       
   365   then show ?thesis
       
   366     unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def continuous_on_const)
       
   367 next
       
   368   case False
       
   369   show ?thesis
       
   370     apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise])
       
   371     apply (rule assms derivative_intros | simp add: False vimage_def)+
       
   372     using real_vector_affinity_eq [OF False, where c=c, unfolded scaleR_conv_of_real]
       
   373     apply simp
       
   374     done
       
   375 qed
       
   376 
       
   377 lemma piecewise_C1_differentiable_cases:
       
   378   fixes c::real
       
   379   assumes "f piecewise_C1_differentiable_on {a..c}"
       
   380           "g piecewise_C1_differentiable_on {c..b}"
       
   381            "a \<le> c" "c \<le> b" "f c = g c"
       
   382   shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_C1_differentiable_on {a..b}"
       
   383 proof -
       
   384   obtain s t where st: "f C1_differentiable_on ({a..c} - s)"
       
   385                        "g C1_differentiable_on ({c..b} - t)"
       
   386                        "finite s" "finite t"
       
   387     using assms
       
   388     by (force simp: piecewise_C1_differentiable_on_def)
       
   389   then have f_diff: "f differentiable_on {a..<c} - s"
       
   390         and g_diff: "g differentiable_on {c<..b} - t"
       
   391     by (simp_all add: C1_differentiable_on_eq differentiable_at_withinI differentiable_on_def)
       
   392   have "continuous_on {a..c} f" "continuous_on {c..b} g"
       
   393     using assms piecewise_C1_differentiable_on_def by auto
       
   394   then have cab: "continuous_on {a..b} (\<lambda>x. if x \<le> c then f x else g x)"
       
   395     using continuous_on_cases [OF closed_real_atLeastAtMost [of a c],
       
   396                                OF closed_real_atLeastAtMost [of c b],
       
   397                                of f g "\<lambda>x. x\<le>c"]  assms
       
   398     by (force simp: ivl_disj_un_two_touch)
       
   399   { fix x
       
   400     assume x: "x \<in> {a..b} - insert c (s \<union> t)"
       
   401     have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x" (is "?diff_fg")
       
   402     proof (cases x c rule: le_cases)
       
   403       case le show ?diff_fg
       
   404         apply (rule differentiable_transform_within [where f=f and d = "dist x c"])
       
   405         using x dist_real_def le st by (auto simp: C1_differentiable_on_eq)
       
   406     next
       
   407       case ge show ?diff_fg
       
   408         apply (rule differentiable_transform_within [where f=g and d = "dist x c"])
       
   409         using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq)
       
   410     qed
       
   411   }
       
   412   then have "(\<forall>x \<in> {a..b} - insert c (s \<union> t). (\<lambda>x. if x \<le> c then f x else g x) differentiable at x)"
       
   413     by auto
       
   414   moreover
       
   415   { assume fcon: "continuous_on ({a<..<c} - s) (\<lambda>x. vector_derivative f (at x))"
       
   416        and gcon: "continuous_on ({c<..<b} - t) (\<lambda>x. vector_derivative g (at x))"
       
   417     have "open ({a<..<c} - s)"  "open ({c<..<b} - t)"
       
   418       using st by (simp_all add: open_Diff finite_imp_closed)
       
   419     moreover have "continuous_on ({a<..<c} - s) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
       
   420       apply (rule continuous_on_eq [OF fcon])
       
   421       apply (simp add:)
       
   422       apply (rule vector_derivative_at [symmetric])
       
   423       apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within)
       
   424       apply (simp_all add: dist_norm vector_derivative_works [symmetric])
       
   425       apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff linorder_not_le order_less_irrefl st(1))
       
   426       apply auto
       
   427       done
       
   428     moreover have "continuous_on ({c<..<b} - t) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
       
   429       apply (rule continuous_on_eq [OF gcon])
       
   430       apply (simp add:)
       
   431       apply (rule vector_derivative_at [symmetric])
       
   432       apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within)
       
   433       apply (simp_all add: dist_norm vector_derivative_works [symmetric])
       
   434       apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff less_irrefl not_le st(2))
       
   435       apply auto
       
   436       done
       
   437     ultimately have "continuous_on ({a<..<b} - insert c (s \<union> t))
       
   438         (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
       
   439       apply (rule continuous_on_subset [OF continuous_on_open_Un], auto)
       
   440       done
       
   441   } note * = this
       
   442   have "continuous_on ({a<..<b} - insert c (s \<union> t)) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
       
   443     using st
       
   444     by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *)
       
   445   ultimately have "\<exists>s. finite s \<and> ((\<lambda>x. if x \<le> c then f x else g x) C1_differentiable_on {a..b} - s)"
       
   446     apply (rule_tac x="{a,b,c} \<union> s \<union> t" in exI)
       
   447     using st  by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset)
       
   448   with cab show ?thesis
       
   449     by (simp add: piecewise_C1_differentiable_on_def)
       
   450 qed
       
   451 
       
   452 lemma piecewise_C1_differentiable_neg:
       
   453     "f piecewise_C1_differentiable_on s \<Longrightarrow> (\<lambda>x. -(f x)) piecewise_C1_differentiable_on s"
       
   454   unfolding piecewise_C1_differentiable_on_def
       
   455   by (auto intro!: continuous_on_minus C1_differentiable_on_minus)
       
   456 
       
   457 lemma piecewise_C1_differentiable_add:
       
   458   assumes "f piecewise_C1_differentiable_on i"
       
   459           "g piecewise_C1_differentiable_on i"
       
   460     shows "(\<lambda>x. f x + g x) piecewise_C1_differentiable_on i"
       
   461 proof -
       
   462   obtain s t where st: "finite s" "finite t"
       
   463                        "f C1_differentiable_on (i-s)"
       
   464                        "g C1_differentiable_on (i-t)"
       
   465     using assms by (auto simp: piecewise_C1_differentiable_on_def)
       
   466   then have "finite (s \<union> t) \<and> (\<lambda>x. f x + g x) C1_differentiable_on i - (s \<union> t)"
       
   467     by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset)
       
   468   moreover have "continuous_on i f" "continuous_on i g"
       
   469     using assms piecewise_C1_differentiable_on_def by auto
       
   470   ultimately show ?thesis
       
   471     by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add)
       
   472 qed
       
   473 
       
   474 lemma piecewise_C1_differentiable_diff:
       
   475     "\<lbrakk>f piecewise_C1_differentiable_on s;  g piecewise_C1_differentiable_on s\<rbrakk>
       
   476      \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_C1_differentiable_on s"
       
   477   unfolding diff_conv_add_uminus
       
   478   by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg)
       
   479 
       
   480 lemma piecewise_C1_differentiable_D1:
       
   481   fixes g1 :: "real \<Rightarrow> 'a::real_normed_field"
       
   482   assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}"
       
   483     shows "g1 piecewise_C1_differentiable_on {0..1}"
       
   484 proof -
       
   485   obtain s where "finite s"
       
   486              and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
       
   487              and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x"
       
   488     using assms  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
   489   then have g1D: "g1 differentiable at x" if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x
       
   490     apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" in differentiable_transform_within)
       
   491     using that
       
   492     apply (simp_all add: dist_real_def joinpaths_def)
       
   493     apply (rule differentiable_chain_at derivative_intros | force)+
       
   494     done
       
   495   have [simp]: "vector_derivative (g1 \<circ> op * 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)"
       
   496                if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x
       
   497     apply (subst vector_derivative_chain_at)
       
   498     using that
       
   499     apply (rule derivative_eq_intros g1D | simp)+
       
   500     done
       
   501   have "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
       
   502     using co12 by (rule continuous_on_subset) force
       
   503   then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 o op*2) (at x))"
       
   504     apply (rule continuous_on_eq [OF _ vector_derivative_at])
       
   505     apply (rule_tac f="g1 o op*2" and d="dist x (1/2)" in has_vector_derivative_transform_within)
       
   506     apply (simp_all add: dist_norm joinpaths_def vector_derivative_works [symmetric])
       
   507     apply (force intro: g1D differentiable_chain_at)
       
   508     apply auto
       
   509     done
       
   510   have "continuous_on ({0..1} - insert 1 (op * 2 ` s))
       
   511                       ((\<lambda>x. 1/2 * vector_derivative (g1 o op*2) (at x)) o op*(1/2))"
       
   512     apply (rule continuous_intros)+
       
   513     using coDhalf
       
   514     apply (simp add: scaleR_conv_of_real image_set_diff image_image)
       
   515     done
       
   516   then have con_g1: "continuous_on ({0..1} - insert 1 (op * 2 ` s)) (\<lambda>x. vector_derivative g1 (at x))"
       
   517     by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
       
   518   have "continuous_on {0..1} g1"
       
   519     using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast
       
   520   with \<open>finite s\<close> show ?thesis
       
   521     apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
   522     apply (rule_tac x="insert 1 ((op*2)`s)" in exI)
       
   523     apply (simp add: g1D con_g1)
       
   524   done
       
   525 qed
       
   526 
       
   527 lemma piecewise_C1_differentiable_D2:
       
   528   fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
       
   529   assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2"
       
   530     shows "g2 piecewise_C1_differentiable_on {0..1}"
       
   531 proof -
       
   532   obtain s where "finite s"
       
   533              and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
       
   534              and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x"
       
   535     using assms  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
   536   then have g2D: "g2 differentiable at x" if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x
       
   537     apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)" in differentiable_transform_within)
       
   538     using that
       
   539     apply (simp_all add: dist_real_def joinpaths_def)
       
   540     apply (auto simp: dist_real_def joinpaths_def field_simps)
       
   541     apply (rule differentiable_chain_at derivative_intros | force)+
       
   542     apply (drule_tac x= "(x + 1) / 2" in bspec, force simp: divide_simps)
       
   543     apply assumption
       
   544     done
       
   545   have [simp]: "vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)"
       
   546                if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x
       
   547     using that  by (auto simp: vector_derivative_chain_at divide_simps g2D)
       
   548   have "continuous_on ({1/2..1} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
       
   549     using co12 by (rule continuous_on_subset) force
       
   550   then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) s) (\<lambda>x. vector_derivative (g2 o (\<lambda>x. 2*x-1)) (at x))"
       
   551     apply (rule continuous_on_eq [OF _ vector_derivative_at])
       
   552     apply (rule_tac f="g2 o (\<lambda>x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within)
       
   553     apply (auto simp: dist_real_def field_simps joinpaths_def vector_derivative_works [symmetric]
       
   554                 intro!: g2D differentiable_chain_at)
       
   555     done
       
   556   have [simp]: "((\<lambda>x. (x + 1) / 2) ` ({0..1} - insert 0 ((\<lambda>x. 2 * x - 1) ` s))) = ({1/2..1} - insert (1/2) s)"
       
   557     apply (simp add: image_set_diff inj_on_def image_image)
       
   558     apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib)
       
   559     done
       
   560   have "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s))
       
   561                       ((\<lambda>x. 1/2 * vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at x)) o (\<lambda>x. (x+1)/2))"
       
   562     by (rule continuous_intros | simp add:  coDhalf)+
       
   563   then have con_g2: "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)) (\<lambda>x. vector_derivative g2 (at x))"
       
   564     by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
       
   565   have "continuous_on {0..1} g2"
       
   566     using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast
       
   567   with \<open>finite s\<close> show ?thesis
       
   568     apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
   569     apply (rule_tac x="insert 0 ((\<lambda>x. 2 * x - 1) ` s)" in exI)
       
   570     apply (simp add: g2D con_g2)
       
   571   done
       
   572 qed
       
   573 
       
   574 subsection \<open>Valid paths, and their start and finish\<close>
       
   575 
       
   576 lemma Diff_Un_eq: "A - (B \<union> C) = A - B - C"
       
   577   by blast
       
   578 
       
   579 definition valid_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
       
   580   where "valid_path f \<equiv> f piecewise_C1_differentiable_on {0..1::real}"
       
   581 
       
   582 definition closed_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
       
   583   where "closed_path g \<equiv> g 0 = g 1"
       
   584 
       
   585 subsubsection\<open>In particular, all results for paths apply\<close>
       
   586 
       
   587 lemma valid_path_imp_path: "valid_path g \<Longrightarrow> path g"
       
   588 by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def)
       
   589 
       
   590 lemma connected_valid_path_image: "valid_path g \<Longrightarrow> connected(path_image g)"
       
   591   by (metis connected_path_image valid_path_imp_path)
       
   592 
       
   593 lemma compact_valid_path_image: "valid_path g \<Longrightarrow> compact(path_image g)"
       
   594   by (metis compact_path_image valid_path_imp_path)
       
   595 
       
   596 lemma bounded_valid_path_image: "valid_path g \<Longrightarrow> bounded(path_image g)"
       
   597   by (metis bounded_path_image valid_path_imp_path)
       
   598 
       
   599 lemma closed_valid_path_image: "valid_path g \<Longrightarrow> closed(path_image g)"
       
   600   by (metis closed_path_image valid_path_imp_path)
       
   601 
       
   602 proposition valid_path_compose:
       
   603   assumes "valid_path g"
       
   604       and der: "\<And>x. x \<in> path_image g \<Longrightarrow> \<exists>f'. (f has_field_derivative f') (at x)"
       
   605       and con: "continuous_on (path_image g) (deriv f)"
       
   606     shows "valid_path (f o g)"
       
   607 proof -
       
   608   obtain s where "finite s" and g_diff: "g C1_differentiable_on {0..1} - s"
       
   609     using \<open>valid_path g\<close> unfolding valid_path_def piecewise_C1_differentiable_on_def by auto
       
   610   have "f \<circ> g differentiable at t" when "t\<in>{0..1} - s" for t
       
   611     proof (rule differentiable_chain_at)
       
   612       show "g differentiable at t" using \<open>valid_path g\<close>
       
   613         by (meson C1_differentiable_on_eq \<open>g C1_differentiable_on {0..1} - s\<close> that)
       
   614     next
       
   615       have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis
       
   616       then obtain f' where "(f has_field_derivative f') (at (g t))"
       
   617         using der by auto
       
   618       then have " (f has_derivative op * f') (at (g t))"
       
   619         using has_field_derivative_imp_has_derivative[of f f' "at (g t)"] by auto
       
   620       then show "f differentiable at (g t)" using differentiableI by auto
       
   621     qed
       
   622   moreover have "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (f \<circ> g) (at x))"
       
   623     proof (rule continuous_on_eq [where f = "\<lambda>x. vector_derivative g (at x) * deriv f (g x)"],
       
   624         rule continuous_intros)
       
   625       show "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative g (at x))"
       
   626         using g_diff C1_differentiable_on_eq by auto
       
   627     next
       
   628       have "continuous_on {0..1} (\<lambda>x. deriv f (g x))"
       
   629         using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def]
       
   630           \<open>valid_path g\<close> piecewise_C1_differentiable_on_def valid_path_def
       
   631         by blast
       
   632       then show "continuous_on ({0..1} - s) (\<lambda>x. deriv f (g x))"
       
   633         using continuous_on_subset by blast
       
   634     next
       
   635       show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \<circ> g) (at t)"
       
   636           when "t \<in> {0..1} - s" for t
       
   637         proof (rule vector_derivative_chain_at_general[symmetric])
       
   638           show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that)
       
   639         next
       
   640           have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis
       
   641           then obtain f' where "(f has_field_derivative f') (at (g t))"
       
   642             using der by auto
       
   643           then show "\<exists>g'. (f has_field_derivative g') (at (g t))" by auto
       
   644         qed
       
   645     qed
       
   646   ultimately have "f o g C1_differentiable_on {0..1} - s"
       
   647     using C1_differentiable_on_eq by blast
       
   648   moreover have "path (f o g)"
       
   649     proof -
       
   650       have "isCont f x" when "x\<in>path_image g" for x
       
   651         proof -
       
   652           obtain f' where "(f has_field_derivative f') (at x)"
       
   653             using der[rule_format] \<open>x\<in>path_image g\<close> by auto
       
   654           thus ?thesis using DERIV_isCont by auto
       
   655         qed
       
   656       then have "continuous_on (path_image g) f" using continuous_at_imp_continuous_on by auto
       
   657       then show ?thesis using path_continuous_image \<open>valid_path g\<close> valid_path_imp_path by auto
       
   658     qed
       
   659   ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def
       
   660     using \<open>finite s\<close> by auto
       
   661 qed
       
   662 
       
   663 
       
   664 subsection\<open>Contour Integrals along a path\<close>
       
   665 
       
   666 text\<open>This definition is for complex numbers only, and does not generalise to line integrals in a vector field\<close>
       
   667 
       
   668 text\<open>piecewise differentiable function on [0,1]\<close>
       
   669 
       
   670 definition has_contour_integral :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool"
       
   671            (infixr "has'_contour'_integral" 50)
       
   672   where "(f has_contour_integral i) g \<equiv>
       
   673            ((\<lambda>x. f(g x) * vector_derivative g (at x within {0..1}))
       
   674             has_integral i) {0..1}"
       
   675 
       
   676 definition contour_integrable_on
       
   677            (infixr "contour'_integrable'_on" 50)
       
   678   where "f contour_integrable_on g \<equiv> \<exists>i. (f has_contour_integral i) g"
       
   679 
       
   680 definition contour_integral
       
   681   where "contour_integral g f \<equiv> @i. (f has_contour_integral i) g \<or> ~ f contour_integrable_on g \<and> i=0"
       
   682 
       
   683 lemma not_integrable_contour_integral: "~ f contour_integrable_on g \<Longrightarrow> contour_integral g f = 0"
       
   684   unfolding contour_integrable_on_def contour_integral_def by blast
       
   685 
       
   686 lemma contour_integral_unique: "(f has_contour_integral i) g \<Longrightarrow> contour_integral g f = i"
       
   687   apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def)
       
   688   using has_integral_unique by blast
       
   689 
       
   690 corollary has_contour_integral_eqpath:
       
   691      "\<lbrakk>(f has_contour_integral y) p; f contour_integrable_on \<gamma>;
       
   692        contour_integral p f = contour_integral \<gamma> f\<rbrakk>
       
   693       \<Longrightarrow> (f has_contour_integral y) \<gamma>"
       
   694 using contour_integrable_on_def contour_integral_unique by auto
       
   695 
       
   696 lemma has_contour_integral_integral:
       
   697     "f contour_integrable_on i \<Longrightarrow> (f has_contour_integral (contour_integral i f)) i"
       
   698   by (metis contour_integral_unique contour_integrable_on_def)
       
   699 
       
   700 lemma has_contour_integral_unique:
       
   701     "(f has_contour_integral i) g \<Longrightarrow> (f has_contour_integral j) g \<Longrightarrow> i = j"
       
   702   using has_integral_unique
       
   703   by (auto simp: has_contour_integral_def)
       
   704 
       
   705 lemma has_contour_integral_integrable: "(f has_contour_integral i) g \<Longrightarrow> f contour_integrable_on g"
       
   706   using contour_integrable_on_def by blast
       
   707 
       
   708 (* Show that we can forget about the localized derivative.*)
       
   709 
       
   710 lemma vector_derivative_within_interior:
       
   711      "\<lbrakk>x \<in> interior s; NO_MATCH UNIV s\<rbrakk>
       
   712       \<Longrightarrow> vector_derivative f (at x within s) = vector_derivative f (at x)"
       
   713   apply (simp add: vector_derivative_def has_vector_derivative_def has_derivative_def netlimit_within_interior)
       
   714   apply (subst lim_within_interior, auto)
       
   715   done
       
   716 
       
   717 lemma has_integral_localized_vector_derivative:
       
   718     "((\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \<longleftrightarrow>
       
   719      ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}"
       
   720 proof -
       
   721   have "{a..b} - {a,b} = interior {a..b}"
       
   722     by (simp add: atLeastAtMost_diff_ends)
       
   723   show ?thesis
       
   724     apply (rule has_integral_spike_eq [of "{a,b}"])
       
   725     apply (auto simp: vector_derivative_within_interior)
       
   726     done
       
   727 qed
       
   728 
       
   729 lemma integrable_on_localized_vector_derivative:
       
   730     "(\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \<longleftrightarrow>
       
   731      (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}"
       
   732   by (simp add: integrable_on_def has_integral_localized_vector_derivative)
       
   733 
       
   734 lemma has_contour_integral:
       
   735      "(f has_contour_integral i) g \<longleftrightarrow>
       
   736       ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}"
       
   737   by (simp add: has_integral_localized_vector_derivative has_contour_integral_def)
       
   738 
       
   739 lemma contour_integrable_on:
       
   740      "f contour_integrable_on g \<longleftrightarrow>
       
   741       (\<lambda>t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}"
       
   742   by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def)
       
   743 
       
   744 subsection\<open>Reversing a path\<close>
       
   745 
       
   746 lemma valid_path_imp_reverse:
       
   747   assumes "valid_path g"
       
   748     shows "valid_path(reversepath g)"
       
   749 proof -
       
   750   obtain s where "finite s" "g C1_differentiable_on ({0..1} - s)"
       
   751     using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def)
       
   752   then have "finite (op - 1 ` s)" "(reversepath g C1_differentiable_on ({0..1} - op - 1 ` s))"
       
   753     apply (auto simp: reversepath_def)
       
   754     apply (rule C1_differentiable_compose [of "\<lambda>x::real. 1-x" _ g, unfolded o_def])
       
   755     apply (auto simp: C1_differentiable_on_eq)
       
   756     apply (rule continuous_intros, force)
       
   757     apply (force elim!: continuous_on_subset)
       
   758     apply (simp add: finite_vimageI inj_on_def)
       
   759     done
       
   760   then show ?thesis using assms
       
   761     by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric])
       
   762 qed
       
   763 
       
   764 lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \<longleftrightarrow> valid_path g"
       
   765   using valid_path_imp_reverse by force
       
   766 
       
   767 lemma has_contour_integral_reversepath:
       
   768   assumes "valid_path g" "(f has_contour_integral i) g"
       
   769     shows "(f has_contour_integral (-i)) (reversepath g)"
       
   770 proof -
       
   771   { fix s x
       
   772     assume xs: "g C1_differentiable_on ({0..1} - s)" "x \<notin> op - 1 ` s" "0 \<le> x" "x \<le> 1"
       
   773       have "vector_derivative (\<lambda>x. g (1 - x)) (at x within {0..1}) =
       
   774             - vector_derivative g (at (1 - x) within {0..1})"
       
   775       proof -
       
   776         obtain f' where f': "(g has_vector_derivative f') (at (1 - x))"
       
   777           using xs
       
   778           by (force simp: has_vector_derivative_def C1_differentiable_on_def)
       
   779         have "(g o (\<lambda>x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)"
       
   780           apply (rule vector_diff_chain_within)
       
   781           apply (intro vector_diff_chain_within derivative_eq_intros | simp)+
       
   782           apply (rule has_vector_derivative_at_within [OF f'])
       
   783           done
       
   784         then have mf': "((\<lambda>x. g (1 - x)) has_vector_derivative -f') (at x)"
       
   785           by (simp add: o_def)
       
   786         show ?thesis
       
   787           using xs
       
   788           by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f'])
       
   789       qed
       
   790   } note * = this
       
   791   have 01: "{0..1::real} = cbox 0 1"
       
   792     by simp
       
   793   show ?thesis using assms
       
   794     apply (auto simp: has_contour_integral_def)
       
   795     apply (drule has_integral_affinity01 [where m= "-1" and c=1])
       
   796     apply (auto simp: reversepath_def valid_path_def piecewise_C1_differentiable_on_def)
       
   797     apply (drule has_integral_neg)
       
   798     apply (rule_tac s = "(\<lambda>x. 1 - x) ` s" in has_integral_spike_finite)
       
   799     apply (auto simp: *)
       
   800     done
       
   801 qed
       
   802 
       
   803 lemma contour_integrable_reversepath:
       
   804     "valid_path g \<Longrightarrow> f contour_integrable_on g \<Longrightarrow> f contour_integrable_on (reversepath g)"
       
   805   using has_contour_integral_reversepath contour_integrable_on_def by blast
       
   806 
       
   807 lemma contour_integrable_reversepath_eq:
       
   808     "valid_path g \<Longrightarrow> (f contour_integrable_on (reversepath g) \<longleftrightarrow> f contour_integrable_on g)"
       
   809   using contour_integrable_reversepath valid_path_reversepath by fastforce
       
   810 
       
   811 lemma contour_integral_reversepath:
       
   812   assumes "valid_path g"
       
   813     shows "contour_integral (reversepath g) f = - (contour_integral g f)"
       
   814 proof (cases "f contour_integrable_on g")
       
   815   case True then show ?thesis
       
   816     by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath)
       
   817 next
       
   818   case False then have "~ f contour_integrable_on (reversepath g)"
       
   819     by (simp add: assms contour_integrable_reversepath_eq)
       
   820   with False show ?thesis by (simp add: not_integrable_contour_integral)
       
   821 qed
       
   822 
       
   823 
       
   824 subsection\<open>Joining two paths together\<close>
       
   825 
       
   826 lemma valid_path_join:
       
   827   assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2"
       
   828     shows "valid_path(g1 +++ g2)"
       
   829 proof -
       
   830   have "g1 1 = g2 0"
       
   831     using assms by (auto simp: pathfinish_def pathstart_def)
       
   832   moreover have "(g1 o (\<lambda>x. 2*x)) piecewise_C1_differentiable_on {0..1/2}"
       
   833     apply (rule piecewise_C1_differentiable_compose)
       
   834     using assms
       
   835     apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths)
       
   836     apply (rule continuous_intros | simp)+
       
   837     apply (force intro: finite_vimageI [where h = "op*2"] inj_onI)
       
   838     done
       
   839   moreover have "(g2 o (\<lambda>x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}"
       
   840     apply (rule piecewise_C1_differentiable_compose)
       
   841     using assms unfolding valid_path_def piecewise_C1_differentiable_on_def
       
   842     by (auto intro!: continuous_intros finite_vimageI [where h = "(\<lambda>x. 2*x - 1)"] inj_onI
       
   843              simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths)
       
   844   ultimately show ?thesis
       
   845     apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def)
       
   846     apply (rule piecewise_C1_differentiable_cases)
       
   847     apply (auto simp: o_def)
       
   848     done
       
   849 qed
       
   850 
       
   851 lemma valid_path_join_D1:
       
   852   fixes g1 :: "real \<Rightarrow> 'a::real_normed_field"
       
   853   shows "valid_path (g1 +++ g2) \<Longrightarrow> valid_path g1"
       
   854   unfolding valid_path_def
       
   855   by (rule piecewise_C1_differentiable_D1)
       
   856 
       
   857 lemma valid_path_join_D2:
       
   858   fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
       
   859   shows "\<lbrakk>valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> valid_path g2"
       
   860   unfolding valid_path_def
       
   861   by (rule piecewise_C1_differentiable_D2)
       
   862 
       
   863 lemma valid_path_join_eq [simp]:
       
   864   fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
       
   865   shows "pathfinish g1 = pathstart g2 \<Longrightarrow> (valid_path(g1 +++ g2) \<longleftrightarrow> valid_path g1 \<and> valid_path g2)"
       
   866   using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast
       
   867 
       
   868 lemma has_contour_integral_join:
       
   869   assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2"
       
   870           "valid_path g1" "valid_path g2"
       
   871     shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)"
       
   872 proof -
       
   873   obtain s1 s2
       
   874     where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x"
       
   875       and s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x"
       
   876     using assms
       
   877     by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
   878   have 1: "((\<lambda>x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}"
       
   879    and 2: "((\<lambda>x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}"
       
   880     using assms
       
   881     by (auto simp: has_contour_integral)
       
   882   have i1: "((\<lambda>x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}"
       
   883    and i2: "((\<lambda>x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}"
       
   884     using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]]
       
   885           has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]]
       
   886     by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac)
       
   887   have g1: "\<lbrakk>0 \<le> z; z*2 < 1; z*2 \<notin> s1\<rbrakk> \<Longrightarrow>
       
   888             vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) =
       
   889             2 *\<^sub>R vector_derivative g1 (at (z*2))" for z
       
   890     apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g1(2*x))" and d = "\<bar>z - 1/2\<bar>"]])
       
   891     apply (simp_all add: dist_real_def abs_if split: if_split_asm)
       
   892     apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x" 2 _ g1, simplified o_def])
       
   893     apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
       
   894     using s1
       
   895     apply (auto simp: algebra_simps vector_derivative_works)
       
   896     done
       
   897   have g2: "\<lbrakk>1 < z*2; z \<le> 1; z*2 - 1 \<notin> s2\<rbrakk> \<Longrightarrow>
       
   898             vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) =
       
   899             2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z
       
   900     apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g2 (2*x - 1))" and d = "\<bar>z - 1/2\<bar>"]])
       
   901     apply (simp_all add: dist_real_def abs_if split: if_split_asm)
       
   902     apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x - 1" 2 _ g2, simplified o_def])
       
   903     apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
       
   904     using s2
       
   905     apply (auto simp: algebra_simps vector_derivative_works)
       
   906     done
       
   907   have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}"
       
   908     apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) (op*2 -` s1)"])
       
   909     using s1
       
   910     apply (force intro: finite_vimageI [where h = "op*2"] inj_onI)
       
   911     apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1)
       
   912     done
       
   913   moreover have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}"
       
   914     apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\<lambda>x. 2*x-1) -` s2)"])
       
   915     using s2
       
   916     apply (force intro: finite_vimageI [where h = "\<lambda>x. 2*x-1"] inj_onI)
       
   917     apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2)
       
   918     done
       
   919   ultimately
       
   920   show ?thesis
       
   921     apply (simp add: has_contour_integral)
       
   922     apply (rule has_integral_combine [where c = "1/2"], auto)
       
   923     done
       
   924 qed
       
   925 
       
   926 lemma contour_integrable_joinI:
       
   927   assumes "f contour_integrable_on g1" "f contour_integrable_on g2"
       
   928           "valid_path g1" "valid_path g2"
       
   929     shows "f contour_integrable_on (g1 +++ g2)"
       
   930   using assms
       
   931   by (meson has_contour_integral_join contour_integrable_on_def)
       
   932 
       
   933 lemma contour_integrable_joinD1:
       
   934   assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1"
       
   935     shows "f contour_integrable_on g1"
       
   936 proof -
       
   937   obtain s1
       
   938     where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x"
       
   939     using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
   940   have "(\<lambda>x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}"
       
   941     using assms
       
   942     apply (auto simp: contour_integrable_on)
       
   943     apply (drule integrable_on_subcbox [where a=0 and b="1/2"])
       
   944     apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified])
       
   945     done
       
   946   then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}"
       
   947     by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real)
       
   948   have g1: "\<lbrakk>0 < z; z < 1; z \<notin> s1\<rbrakk> \<Longrightarrow>
       
   949             vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) =
       
   950             2 *\<^sub>R vector_derivative g1 (at z)"  for z
       
   951     apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g1(2*x))" and d = "\<bar>(z-1)/2\<bar>"]])
       
   952     apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm)
       
   953     apply (rule vector_diff_chain_at [of "\<lambda>x. x*2" 2 _ g1, simplified o_def])
       
   954     using s1
       
   955     apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
       
   956     done
       
   957   show ?thesis
       
   958     using s1
       
   959     apply (auto simp: contour_integrable_on)
       
   960     apply (rule integrable_spike_finite [of "{0,1} \<union> s1", OF _ _ *])
       
   961     apply (auto simp: joinpaths_def scaleR_conv_of_real g1)
       
   962     done
       
   963 qed
       
   964 
       
   965 lemma contour_integrable_joinD2:
       
   966   assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2"
       
   967     shows "f contour_integrable_on g2"
       
   968 proof -
       
   969   obtain s2
       
   970     where s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x"
       
   971     using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
   972   have "(\<lambda>x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}"
       
   973     using assms
       
   974     apply (auto simp: contour_integrable_on)
       
   975     apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto)
       
   976     apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified])
       
   977     apply (simp add: image_affinity_atLeastAtMost_diff)
       
   978     done
       
   979   then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2)))
       
   980                 integrable_on {0..1}"
       
   981     by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real)
       
   982   have g2: "\<lbrakk>0 < z; z < 1; z \<notin> s2\<rbrakk> \<Longrightarrow>
       
   983             vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) =
       
   984             2 *\<^sub>R vector_derivative g2 (at z)" for z
       
   985     apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g2(2*x-1))" and d = "\<bar>z/2\<bar>"]])
       
   986     apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm)
       
   987     apply (rule vector_diff_chain_at [of "\<lambda>x. x*2-1" 2 _ g2, simplified o_def])
       
   988     using s2
       
   989     apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left
       
   990                       vector_derivative_works add_divide_distrib)
       
   991     done
       
   992   show ?thesis
       
   993     using s2
       
   994     apply (auto simp: contour_integrable_on)
       
   995     apply (rule integrable_spike_finite [of "{0,1} \<union> s2", OF _ _ *])
       
   996     apply (auto simp: joinpaths_def scaleR_conv_of_real g2)
       
   997     done
       
   998 qed
       
   999 
       
  1000 lemma contour_integrable_join [simp]:
       
  1001   shows
       
  1002     "\<lbrakk>valid_path g1; valid_path g2\<rbrakk>
       
  1003      \<Longrightarrow> f contour_integrable_on (g1 +++ g2) \<longleftrightarrow> f contour_integrable_on g1 \<and> f contour_integrable_on g2"
       
  1004 using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast
       
  1005 
       
  1006 lemma contour_integral_join [simp]:
       
  1007   shows
       
  1008     "\<lbrakk>f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\<rbrakk>
       
  1009         \<Longrightarrow> contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f"
       
  1010   by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique)
       
  1011 
       
  1012 
       
  1013 subsection\<open>Shifting the starting point of a (closed) path\<close>
       
  1014 
       
  1015 lemma shiftpath_alt_def: "shiftpath a f = (\<lambda>x. if x \<le> 1-a then f (a + x) else f (a + x - 1))"
       
  1016   by (auto simp: shiftpath_def)
       
  1017 
       
  1018 lemma valid_path_shiftpath [intro]:
       
  1019   assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
       
  1020     shows "valid_path(shiftpath a g)"
       
  1021   using assms
       
  1022   apply (auto simp: valid_path_def shiftpath_alt_def)
       
  1023   apply (rule piecewise_C1_differentiable_cases)
       
  1024   apply (auto simp: algebra_simps)
       
  1025   apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one])
       
  1026   apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
       
  1027   apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps])
       
  1028   apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
       
  1029   done
       
  1030 
       
  1031 lemma has_contour_integral_shiftpath:
       
  1032   assumes f: "(f has_contour_integral i) g" "valid_path g"
       
  1033       and a: "a \<in> {0..1}"
       
  1034     shows "(f has_contour_integral i) (shiftpath a g)"
       
  1035 proof -
       
  1036   obtain s
       
  1037     where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x"
       
  1038     using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
  1039   have *: "((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}"
       
  1040     using assms by (auto simp: has_contour_integral)
       
  1041   then have i: "i = integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)) +
       
  1042                     integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x))"
       
  1043     apply (rule has_integral_unique)
       
  1044     apply (subst add.commute)
       
  1045     apply (subst Integration.integral_combine)
       
  1046     using assms * integral_unique by auto
       
  1047   { fix x
       
  1048     have "0 \<le> x \<Longrightarrow> x + a < 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a) ` s \<Longrightarrow>
       
  1049          vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))"
       
  1050       unfolding shiftpath_def
       
  1051       apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g(a+x))" and d = "dist(1-a) x"]])
       
  1052         apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm)
       
  1053       apply (rule vector_diff_chain_at [of "\<lambda>x. x+a" 1 _ g, simplified o_def scaleR_one])
       
  1054        apply (intro derivative_eq_intros | simp)+
       
  1055       using g
       
  1056        apply (drule_tac x="x+a" in bspec)
       
  1057       using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute)
       
  1058       done
       
  1059   } note vd1 = this
       
  1060   { fix x
       
  1061     have "1 < x + a \<Longrightarrow> x \<le> 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a + 1) ` s \<Longrightarrow>
       
  1062           vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))"
       
  1063       unfolding shiftpath_def
       
  1064       apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g(a+x-1))" and d = "dist (1-a) x"]])
       
  1065         apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm)
       
  1066       apply (rule vector_diff_chain_at [of "\<lambda>x. x+a-1" 1 _ g, simplified o_def scaleR_one])
       
  1067        apply (intro derivative_eq_intros | simp)+
       
  1068       using g
       
  1069       apply (drule_tac x="x+a-1" in bspec)
       
  1070       using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute)
       
  1071       done
       
  1072   } note vd2 = this
       
  1073   have va1: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})"
       
  1074     using * a   by (fastforce intro: integrable_subinterval_real)
       
  1075   have v0a: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})"
       
  1076     apply (rule integrable_subinterval_real)
       
  1077     using * a by auto
       
  1078   have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x))
       
  1079         has_integral  integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)))  {0..1 - a}"
       
  1080     apply (rule has_integral_spike_finite
       
  1081              [where s = "{1-a} \<union> (\<lambda>x. x-a) ` s" and f = "\<lambda>x. f(g(a+x)) * vector_derivative g (at(a+x))"])
       
  1082       using s apply blast
       
  1083      using a apply (auto simp: algebra_simps vd1)
       
  1084      apply (force simp: shiftpath_def add.commute)
       
  1085     using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]]
       
  1086     apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute)
       
  1087     done
       
  1088   moreover
       
  1089   have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x))
       
  1090         has_integral  integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x)))  {1 - a..1}"
       
  1091     apply (rule has_integral_spike_finite
       
  1092              [where s = "{1-a} \<union> (\<lambda>x. x-a+1) ` s" and f = "\<lambda>x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"])
       
  1093       using s apply blast
       
  1094      using a apply (auto simp: algebra_simps vd2)
       
  1095      apply (force simp: shiftpath_def add.commute)
       
  1096     using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]]
       
  1097     apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified])
       
  1098     apply (simp add: algebra_simps)
       
  1099     done
       
  1100   ultimately show ?thesis
       
  1101     using a
       
  1102     by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"])
       
  1103 qed
       
  1104 
       
  1105 lemma has_contour_integral_shiftpath_D:
       
  1106   assumes "(f has_contour_integral i) (shiftpath a g)"
       
  1107           "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
       
  1108     shows "(f has_contour_integral i) g"
       
  1109 proof -
       
  1110   obtain s
       
  1111     where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x"
       
  1112     using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
       
  1113   { fix x
       
  1114     assume x: "0 < x" "x < 1" "x \<notin> s"
       
  1115     then have gx: "g differentiable at x"
       
  1116       using g by auto
       
  1117     have "vector_derivative g (at x within {0..1}) =
       
  1118           vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})"
       
  1119       apply (rule vector_derivative_at_within_ivl
       
  1120                   [OF has_vector_derivative_transform_within_open
       
  1121                       [where f = "(shiftpath (1 - a) (shiftpath a g))" and s = "{0<..<1}-s"]])
       
  1122       using s g assms x
       
  1123       apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath
       
  1124                         vector_derivative_within_interior vector_derivative_works [symmetric])
       
  1125       apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"])
       
  1126       apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm)
       
  1127       done
       
  1128   } note vd = this
       
  1129   have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))"
       
  1130     using assms  by (auto intro!: has_contour_integral_shiftpath)
       
  1131   show ?thesis
       
  1132     apply (simp add: has_contour_integral_def)
       
  1133     apply (rule has_integral_spike_finite [of "{0,1} \<union> s", OF _ _  fi [unfolded has_contour_integral_def]])
       
  1134     using s assms vd
       
  1135     apply (auto simp: Path_Connected.shiftpath_shiftpath)
       
  1136     done
       
  1137 qed
       
  1138 
       
  1139 lemma has_contour_integral_shiftpath_eq:
       
  1140   assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
       
  1141     shows "(f has_contour_integral i) (shiftpath a g) \<longleftrightarrow> (f has_contour_integral i) g"
       
  1142   using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast
       
  1143 
       
  1144 lemma contour_integrable_on_shiftpath_eq:
       
  1145   assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
       
  1146     shows "f contour_integrable_on (shiftpath a g) \<longleftrightarrow> f contour_integrable_on g"
       
  1147 using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto
       
  1148 
       
  1149 lemma contour_integral_shiftpath:
       
  1150   assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
       
  1151     shows "contour_integral (shiftpath a g) f = contour_integral g f"
       
  1152    using assms
       
  1153    by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq)
       
  1154 
       
  1155 
       
  1156 subsection\<open>More about straight-line paths\<close>
       
  1157 
       
  1158 lemma has_vector_derivative_linepath_within:
       
  1159     "(linepath a b has_vector_derivative (b - a)) (at x within s)"
       
  1160 apply (simp add: linepath_def has_vector_derivative_def algebra_simps)
       
  1161 apply (rule derivative_eq_intros | simp)+
       
  1162 done
       
  1163 
       
  1164 lemma vector_derivative_linepath_within:
       
  1165     "x \<in> {0..1} \<Longrightarrow> vector_derivative (linepath a b) (at x within {0..1}) = b - a"
       
  1166   apply (rule vector_derivative_within_closed_interval [of 0 "1::real", simplified])
       
  1167   apply (auto simp: has_vector_derivative_linepath_within)
       
  1168   done
       
  1169 
       
  1170 lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a"
       
  1171   by (simp add: has_vector_derivative_linepath_within vector_derivative_at)
       
  1172 
       
  1173 lemma valid_path_linepath [iff]: "valid_path (linepath a b)"
       
  1174   apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath)
       
  1175   apply (rule_tac x="{}" in exI)
       
  1176   apply (simp add: differentiable_on_def differentiable_def)
       
  1177   using has_vector_derivative_def has_vector_derivative_linepath_within
       
  1178   apply (fastforce simp add: continuous_on_eq_continuous_within)
       
  1179   done
       
  1180 
       
  1181 lemma has_contour_integral_linepath:
       
  1182   shows "(f has_contour_integral i) (linepath a b) \<longleftrightarrow>
       
  1183          ((\<lambda>x. f(linepath a b x) * (b - a)) has_integral i) {0..1}"
       
  1184   by (simp add: has_contour_integral vector_derivative_linepath_at)
       
  1185 
       
  1186 lemma linepath_in_path:
       
  1187   shows "x \<in> {0..1} \<Longrightarrow> linepath a b x \<in> closed_segment a b"
       
  1188   by (auto simp: segment linepath_def)
       
  1189 
       
  1190 lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b"
       
  1191   by (auto simp: segment linepath_def)
       
  1192 
       
  1193 lemma linepath_in_convex_hull:
       
  1194     fixes x::real
       
  1195     assumes a: "a \<in> convex hull s"
       
  1196         and b: "b \<in> convex hull s"
       
  1197         and x: "0\<le>x" "x\<le>1"
       
  1198        shows "linepath a b x \<in> convex hull s"
       
  1199   apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD])
       
  1200   using x
       
  1201   apply (auto simp: linepath_image_01 [symmetric])
       
  1202   done
       
  1203 
       
  1204 lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b"
       
  1205   by (simp add: linepath_def)
       
  1206 
       
  1207 lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0"
       
  1208   by (simp add: linepath_def)
       
  1209 
       
  1210 lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)"
       
  1211   by (simp add: has_contour_integral_linepath)
       
  1212 
       
  1213 lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0"
       
  1214   using has_contour_integral_trivial contour_integral_unique by blast
       
  1215 
       
  1216 
       
  1217 subsection\<open>Relation to subpath construction\<close>
       
  1218 
       
  1219 lemma valid_path_subpath:
       
  1220   fixes g :: "real \<Rightarrow> 'a :: real_normed_vector"
       
  1221   assumes "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
       
  1222     shows "valid_path(subpath u v g)"
       
  1223 proof (cases "v=u")
       
  1224   case True
       
  1225   then show ?thesis
       
  1226     unfolding valid_path_def subpath_def
       
  1227     by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise)
       
  1228 next
       
  1229   case False
       
  1230   have "(g o (\<lambda>x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}"
       
  1231     apply (rule piecewise_C1_differentiable_compose)
       
  1232     apply (simp add: C1_differentiable_imp_piecewise)
       
  1233      apply (simp add: image_affinity_atLeastAtMost)
       
  1234     using assms False
       
  1235     apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset)
       
  1236     apply (subst Int_commute)
       
  1237     apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI)
       
  1238     done
       
  1239   then show ?thesis
       
  1240     by (auto simp: o_def valid_path_def subpath_def)
       
  1241 qed
       
  1242 
       
  1243 lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)"
       
  1244   by (simp add: has_contour_integral subpath_def)
       
  1245 
       
  1246 lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)"
       
  1247   using has_contour_integral_subpath_refl contour_integrable_on_def by blast
       
  1248 
       
  1249 lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0"
       
  1250   by (simp add: has_contour_integral_subpath_refl contour_integral_unique)
       
  1251 
       
  1252 lemma has_contour_integral_subpath:
       
  1253   assumes f: "f contour_integrable_on g" and g: "valid_path g"
       
  1254       and uv: "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
       
  1255     shows "(f has_contour_integral  integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x)))
       
  1256            (subpath u v g)"
       
  1257 proof (cases "v=u")
       
  1258   case True
       
  1259   then show ?thesis
       
  1260     using f   by (simp add: contour_integrable_on_def subpath_def has_contour_integral)
       
  1261 next
       
  1262   case False
       
  1263   obtain s where s: "\<And>x. x \<in> {0..1} - s \<Longrightarrow> g differentiable at x" and fs: "finite s"
       
  1264     using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast
       
  1265   have *: "((\<lambda>x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u)))
       
  1266             has_integral (1 / (v - u)) * integral {u..v} (\<lambda>t. f (g t) * vector_derivative g (at t)))
       
  1267            {0..1}"
       
  1268     using f uv
       
  1269     apply (simp add: contour_integrable_on subpath_def has_contour_integral)
       
  1270     apply (drule integrable_on_subcbox [where a=u and b=v, simplified])
       
  1271     apply (simp_all add: has_integral_integral)
       
  1272     apply (drule has_integral_affinity [where m="v-u" and c=u, simplified])
       
  1273     apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real)
       
  1274     apply (simp add: divide_simps False)
       
  1275     done
       
  1276   { fix x
       
  1277     have "x \<in> {0..1} \<Longrightarrow>
       
  1278            x \<notin> (\<lambda>t. (v-u) *\<^sub>R t + u) -` s \<Longrightarrow>
       
  1279            vector_derivative (\<lambda>x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))"
       
  1280       apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]])
       
  1281       apply (intro derivative_eq_intros | simp)+
       
  1282       apply (cut_tac s [of "(v - u) * x + u"])
       
  1283       using uv mult_left_le [of x "v-u"]
       
  1284       apply (auto simp:  vector_derivative_works)
       
  1285       done
       
  1286   } note vd = this
       
  1287   show ?thesis
       
  1288     apply (cut_tac has_integral_cmul [OF *, where c = "v-u"])
       
  1289     using fs assms
       
  1290     apply (simp add: False subpath_def has_contour_integral)
       
  1291     apply (rule_tac s = "(\<lambda>t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite)
       
  1292     apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real)
       
  1293     done
       
  1294 qed
       
  1295 
       
  1296 lemma contour_integrable_subpath:
       
  1297   assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
       
  1298     shows "f contour_integrable_on (subpath u v g)"
       
  1299   apply (cases u v rule: linorder_class.le_cases)
       
  1300    apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms])
       
  1301   apply (subst reversepath_subpath [symmetric])
       
  1302   apply (rule contour_integrable_reversepath)
       
  1303    using assms apply (blast intro: valid_path_subpath)
       
  1304   apply (simp add: contour_integrable_on_def)
       
  1305   using assms apply (blast intro: has_contour_integral_subpath)
       
  1306   done
       
  1307 
       
  1308 lemma has_integral_integrable_integral: "(f has_integral i) s \<longleftrightarrow> f integrable_on s \<and> integral s f = i"
       
  1309   by blast
       
  1310 
       
  1311 lemma has_integral_contour_integral_subpath:
       
  1312   assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
       
  1313     shows "(((\<lambda>x. f(g x) * vector_derivative g (at x)))
       
  1314             has_integral  contour_integral (subpath u v g) f) {u..v}"
       
  1315   using assms
       
  1316   apply (auto simp: has_integral_integrable_integral)
       
  1317   apply (rule integrable_on_subcbox [where a=u and b=v and s = "{0..1}", simplified])
       
  1318   apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on)
       
  1319   done
       
  1320 
       
  1321 lemma contour_integral_subcontour_integral:
       
  1322   assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
       
  1323     shows "contour_integral (subpath u v g) f =
       
  1324            integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x))"
       
  1325   using assms has_contour_integral_subpath contour_integral_unique by blast
       
  1326 
       
  1327 lemma contour_integral_subpath_combine_less:
       
  1328   assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}"
       
  1329           "u<v" "v<w"
       
  1330     shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f =
       
  1331            contour_integral (subpath u w g) f"
       
  1332   using assms apply (auto simp: contour_integral_subcontour_integral)
       
  1333   apply (rule integral_combine, auto)
       
  1334   apply (rule integrable_on_subcbox [where a=u and b=w and s = "{0..1}", simplified])
       
  1335   apply (auto simp: contour_integrable_on)
       
  1336   done
       
  1337 
       
  1338 lemma contour_integral_subpath_combine:
       
  1339   assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}"
       
  1340     shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f =
       
  1341            contour_integral (subpath u w g) f"
       
  1342 proof (cases "u\<noteq>v \<and> v\<noteq>w \<and> u\<noteq>w")
       
  1343   case True
       
  1344     have *: "subpath v u g = reversepath(subpath u v g) \<and>
       
  1345              subpath w u g = reversepath(subpath u w g) \<and>
       
  1346              subpath w v g = reversepath(subpath v w g)"
       
  1347       by (auto simp: reversepath_subpath)
       
  1348     have "u < v \<and> v < w \<or>
       
  1349           u < w \<and> w < v \<or>
       
  1350           v < u \<and> u < w \<or>
       
  1351           v < w \<and> w < u \<or>
       
  1352           w < u \<and> u < v \<or>
       
  1353           w < v \<and> v < u"
       
  1354       using True assms by linarith
       
  1355     with assms show ?thesis
       
  1356       using contour_integral_subpath_combine_less [of f g u v w]
       
  1357             contour_integral_subpath_combine_less [of f g u w v]
       
  1358             contour_integral_subpath_combine_less [of f g v u w]
       
  1359             contour_integral_subpath_combine_less [of f g v w u]
       
  1360             contour_integral_subpath_combine_less [of f g w u v]
       
  1361             contour_integral_subpath_combine_less [of f g w v u]
       
  1362       apply simp
       
  1363       apply (elim disjE)
       
  1364       apply (auto simp: * contour_integral_reversepath contour_integrable_subpath
       
  1365                    valid_path_reversepath valid_path_subpath algebra_simps)
       
  1366       done
       
  1367 next
       
  1368   case False
       
  1369   then show ?thesis
       
  1370     apply (auto simp: contour_integral_subpath_refl)
       
  1371     using assms
       
  1372     by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath)
       
  1373 qed
       
  1374 
       
  1375 lemma contour_integral_integral:
       
  1376      "contour_integral g f = integral {0..1} (\<lambda>x. f (g x) * vector_derivative g (at x))"
       
  1377   by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on)
       
  1378 
       
  1379 
       
  1380 text\<open>Cauchy's theorem where there's a primitive\<close>
       
  1381 
       
  1382 lemma contour_integral_primitive_lemma:
       
  1383   fixes f :: "complex \<Rightarrow> complex" and g :: "real \<Rightarrow> complex"
       
  1384   assumes "a \<le> b"
       
  1385       and "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
       
  1386       and "g piecewise_differentiable_on {a..b}"  "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s"
       
  1387     shows "((\<lambda>x. f'(g x) * vector_derivative g (at x within {a..b}))
       
  1388              has_integral (f(g b) - f(g a))) {a..b}"
       
  1389 proof -
       
  1390   obtain k where k: "finite k" "\<forall>x\<in>{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g"
       
  1391     using assms by (auto simp: piecewise_differentiable_on_def)
       
  1392   have cfg: "continuous_on {a..b} (\<lambda>x. f (g x))"
       
  1393     apply (rule continuous_on_compose [OF cg, unfolded o_def])
       
  1394     using assms
       
  1395     apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff)
       
  1396     done
       
  1397   { fix x::real
       
  1398     assume a: "a < x" and b: "x < b" and xk: "x \<notin> k"
       
  1399     then have "g differentiable at x within {a..b}"
       
  1400       using k by (simp add: differentiable_at_withinI)
       
  1401     then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})"
       
  1402       by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real)
       
  1403     then have gdiff: "(g has_derivative (\<lambda>u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})"
       
  1404       by (simp add: has_vector_derivative_def scaleR_conv_of_real)
       
  1405     have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})"
       
  1406       using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def)
       
  1407     then have fdiff: "(f has_derivative op * (f' (g x))) (at (g x) within g ` {a..b})"
       
  1408       by (simp add: has_field_derivative_def)
       
  1409     have "((\<lambda>x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})"
       
  1410       using diff_chain_within [OF gdiff fdiff]
       
  1411       by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac)
       
  1412   } note * = this
       
  1413   show ?thesis
       
  1414     apply (rule fundamental_theorem_of_calculus_interior_strong)
       
  1415     using k assms cfg *
       
  1416     apply (auto simp: at_within_closed_interval)
       
  1417     done
       
  1418 qed
       
  1419 
       
  1420 lemma contour_integral_primitive:
       
  1421   assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
       
  1422       and "valid_path g" "path_image g \<subseteq> s"
       
  1423     shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g"
       
  1424   using assms
       
  1425   apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def)
       
  1426   apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s])
       
  1427   done
       
  1428 
       
  1429 corollary Cauchy_theorem_primitive:
       
  1430   assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
       
  1431       and "valid_path g"  "path_image g \<subseteq> s" "pathfinish g = pathstart g"
       
  1432     shows "(f' has_contour_integral 0) g"
       
  1433   using assms
       
  1434   by (metis diff_self contour_integral_primitive)
       
  1435 
       
  1436 
       
  1437 text\<open>Existence of path integral for continuous function\<close>
       
  1438 lemma contour_integrable_continuous_linepath:
       
  1439   assumes "continuous_on (closed_segment a b) f"
       
  1440   shows "f contour_integrable_on (linepath a b)"
       
  1441 proof -
       
  1442   have "continuous_on {0..1} ((\<lambda>x. f x * (b - a)) o linepath a b)"
       
  1443     apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01)
       
  1444     apply (rule continuous_intros | simp add: assms)+
       
  1445     done
       
  1446   then show ?thesis
       
  1447     apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric])
       
  1448     apply (rule integrable_continuous [of 0 "1::real", simplified])
       
  1449     apply (rule continuous_on_eq [where f = "\<lambda>x. f(linepath a b x)*(b - a)"])
       
  1450     apply (auto simp: vector_derivative_linepath_within)
       
  1451     done
       
  1452 qed
       
  1453 
       
  1454 lemma has_field_der_id: "((\<lambda>x. x\<^sup>2 / 2) has_field_derivative x) (at x)"
       
  1455   by (rule has_derivative_imp_has_field_derivative)
       
  1456      (rule derivative_intros | simp)+
       
  1457 
       
  1458 lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\<lambda>y. y) = (b^2 - a^2)/2"
       
  1459   apply (rule contour_integral_unique)
       
  1460   using contour_integral_primitive [of UNIV "\<lambda>x. x^2/2" "\<lambda>x. x" "linepath a b"]
       
  1461   apply (auto simp: field_simps has_field_der_id)
       
  1462   done
       
  1463 
       
  1464 lemma contour_integrable_on_const [iff]: "(\<lambda>x. c) contour_integrable_on (linepath a b)"
       
  1465   by (simp add: continuous_on_const contour_integrable_continuous_linepath)
       
  1466 
       
  1467 lemma contour_integrable_on_id [iff]: "(\<lambda>x. x) contour_integrable_on (linepath a b)"
       
  1468   by (simp add: continuous_on_id contour_integrable_continuous_linepath)
       
  1469 
       
  1470 
       
  1471 subsection\<open>Arithmetical combining theorems\<close>
       
  1472 
       
  1473 lemma has_contour_integral_neg:
       
  1474     "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. -(f x)) has_contour_integral (-i)) g"
       
  1475   by (simp add: has_integral_neg has_contour_integral_def)
       
  1476 
       
  1477 lemma has_contour_integral_add:
       
  1478     "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk>
       
  1479      \<Longrightarrow> ((\<lambda>x. f1 x + f2 x) has_contour_integral (i1 + i2)) g"
       
  1480   by (simp add: has_integral_add has_contour_integral_def algebra_simps)
       
  1481 
       
  1482 lemma has_contour_integral_diff:
       
  1483   "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk>
       
  1484          \<Longrightarrow> ((\<lambda>x. f1 x - f2 x) has_contour_integral (i1 - i2)) g"
       
  1485   by (simp add: has_integral_sub has_contour_integral_def algebra_simps)
       
  1486 
       
  1487 lemma has_contour_integral_lmul:
       
  1488   "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. c * (f x)) has_contour_integral (c*i)) g"
       
  1489 apply (simp add: has_contour_integral_def)
       
  1490 apply (drule has_integral_mult_right)
       
  1491 apply (simp add: algebra_simps)
       
  1492 done
       
  1493 
       
  1494 lemma has_contour_integral_rmul:
       
  1495   "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. (f x) * c) has_contour_integral (i*c)) g"
       
  1496 apply (drule has_contour_integral_lmul)
       
  1497 apply (simp add: mult.commute)
       
  1498 done
       
  1499 
       
  1500 lemma has_contour_integral_div:
       
  1501   "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. f x/c) has_contour_integral (i/c)) g"
       
  1502   by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul)
       
  1503 
       
  1504 lemma has_contour_integral_eq:
       
  1505     "\<lbrakk>(f has_contour_integral y) p; \<And>x. x \<in> path_image p \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> (g has_contour_integral y) p"
       
  1506 apply (simp add: path_image_def has_contour_integral_def)
       
  1507 by (metis (no_types, lifting) image_eqI has_integral_eq)
       
  1508 
       
  1509 lemma has_contour_integral_bound_linepath:
       
  1510   assumes "(f has_contour_integral i) (linepath a b)"
       
  1511           "0 \<le> B" "\<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B"
       
  1512     shows "norm i \<le> B * norm(b - a)"
       
  1513 proof -
       
  1514   { fix x::real
       
  1515     assume x: "0 \<le> x" "x \<le> 1"
       
  1516   have "norm (f (linepath a b x)) *
       
  1517         norm (vector_derivative (linepath a b) (at x within {0..1})) \<le> B * norm (b - a)"
       
  1518     by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x)
       
  1519   } note * = this
       
  1520   have "norm i \<le> (B * norm (b - a)) * content (cbox 0 (1::real))"
       
  1521     apply (rule has_integral_bound
       
  1522        [of _ "\<lambda>x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"])
       
  1523     using assms * unfolding has_contour_integral_def
       
  1524     apply (auto simp: norm_mult)
       
  1525     done
       
  1526   then show ?thesis
       
  1527     by (auto simp: content_real)
       
  1528 qed
       
  1529 
       
  1530 (*UNUSED
       
  1531 lemma has_contour_integral_bound_linepath_strong:
       
  1532   fixes a :: real and f :: "complex \<Rightarrow> real"
       
  1533   assumes "(f has_contour_integral i) (linepath a b)"
       
  1534           "finite k"
       
  1535           "0 \<le> B" "\<And>x::real. x \<in> closed_segment a b - k \<Longrightarrow> norm(f x) \<le> B"
       
  1536     shows "norm i \<le> B*norm(b - a)"
       
  1537 *)
       
  1538 
       
  1539 lemma has_contour_integral_const_linepath: "((\<lambda>x. c) has_contour_integral c*(b - a))(linepath a b)"
       
  1540   unfolding has_contour_integral_linepath
       
  1541   by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one)
       
  1542 
       
  1543 lemma has_contour_integral_0: "((\<lambda>x. 0) has_contour_integral 0) g"
       
  1544   by (simp add: has_contour_integral_def)
       
  1545 
       
  1546 lemma has_contour_integral_is_0:
       
  1547     "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> (f has_contour_integral 0) g"
       
  1548   by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto
       
  1549 
       
  1550 lemma has_contour_integral_setsum:
       
  1551     "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a has_contour_integral i a) p\<rbrakk>
       
  1552      \<Longrightarrow> ((\<lambda>x. setsum (\<lambda>a. f a x) s) has_contour_integral setsum i s) p"
       
  1553   by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add)
       
  1554 
       
  1555 
       
  1556 subsection \<open>Operations on path integrals\<close>
       
  1557 
       
  1558 lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\<lambda>x. c) = c*(b - a)"
       
  1559   by (rule contour_integral_unique [OF has_contour_integral_const_linepath])
       
  1560 
       
  1561 lemma contour_integral_neg:
       
  1562     "f contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. -(f x)) = -(contour_integral g f)"
       
  1563   by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg)
       
  1564 
       
  1565 lemma contour_integral_add:
       
  1566     "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x + f2 x) =
       
  1567                 contour_integral g f1 + contour_integral g f2"
       
  1568   by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add)
       
  1569 
       
  1570 lemma contour_integral_diff:
       
  1571     "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x - f2 x) =
       
  1572                 contour_integral g f1 - contour_integral g f2"
       
  1573   by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff)
       
  1574 
       
  1575 lemma contour_integral_lmul:
       
  1576   shows "f contour_integrable_on g
       
  1577            \<Longrightarrow> contour_integral g (\<lambda>x. c * f x) = c*contour_integral g f"
       
  1578   by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul)
       
  1579 
       
  1580 lemma contour_integral_rmul:
       
  1581   shows "f contour_integrable_on g
       
  1582         \<Longrightarrow> contour_integral g (\<lambda>x. f x * c) = contour_integral g f * c"
       
  1583   by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul)
       
  1584 
       
  1585 lemma contour_integral_div:
       
  1586   shows "f contour_integrable_on g
       
  1587         \<Longrightarrow> contour_integral g (\<lambda>x. f x / c) = contour_integral g f / c"
       
  1588   by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div)
       
  1589 
       
  1590 lemma contour_integral_eq:
       
  1591     "(\<And>x. x \<in> path_image p \<Longrightarrow> f x = g x) \<Longrightarrow> contour_integral p f = contour_integral p g"
       
  1592   apply (simp add: contour_integral_def)
       
  1593   using has_contour_integral_eq
       
  1594   by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral)
       
  1595 
       
  1596 lemma contour_integral_eq_0:
       
  1597     "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> contour_integral g f = 0"
       
  1598   by (simp add: has_contour_integral_is_0 contour_integral_unique)
       
  1599 
       
  1600 lemma contour_integral_bound_linepath:
       
  1601   shows
       
  1602     "\<lbrakk>f contour_integrable_on (linepath a b);
       
  1603       0 \<le> B; \<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
       
  1604      \<Longrightarrow> norm(contour_integral (linepath a b) f) \<le> B*norm(b - a)"
       
  1605   apply (rule has_contour_integral_bound_linepath [of f])
       
  1606   apply (auto simp: has_contour_integral_integral)
       
  1607   done
       
  1608 
       
  1609 lemma contour_integral_0 [simp]: "contour_integral g (\<lambda>x. 0) = 0"
       
  1610   by (simp add: contour_integral_unique has_contour_integral_0)
       
  1611 
       
  1612 lemma contour_integral_setsum:
       
  1613     "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
       
  1614      \<Longrightarrow> contour_integral p (\<lambda>x. setsum (\<lambda>a. f a x) s) = setsum (\<lambda>a. contour_integral p (f a)) s"
       
  1615   by (auto simp: contour_integral_unique has_contour_integral_setsum has_contour_integral_integral)
       
  1616 
       
  1617 lemma contour_integrable_eq:
       
  1618     "\<lbrakk>f contour_integrable_on p; \<And>x. x \<in> path_image p \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g contour_integrable_on p"
       
  1619   unfolding contour_integrable_on_def
       
  1620   by (metis has_contour_integral_eq)
       
  1621 
       
  1622 
       
  1623 subsection \<open>Arithmetic theorems for path integrability\<close>
       
  1624 
       
  1625 lemma contour_integrable_neg:
       
  1626     "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. -(f x)) contour_integrable_on g"
       
  1627   using has_contour_integral_neg contour_integrable_on_def by blast
       
  1628 
       
  1629 lemma contour_integrable_add:
       
  1630     "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x + f2 x) contour_integrable_on g"
       
  1631   using has_contour_integral_add contour_integrable_on_def
       
  1632   by fastforce
       
  1633 
       
  1634 lemma contour_integrable_diff:
       
  1635     "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x - f2 x) contour_integrable_on g"
       
  1636   using has_contour_integral_diff contour_integrable_on_def
       
  1637   by fastforce
       
  1638 
       
  1639 lemma contour_integrable_lmul:
       
  1640     "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. c * f x) contour_integrable_on g"
       
  1641   using has_contour_integral_lmul contour_integrable_on_def
       
  1642   by fastforce
       
  1643 
       
  1644 lemma contour_integrable_rmul:
       
  1645     "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x * c) contour_integrable_on g"
       
  1646   using has_contour_integral_rmul contour_integrable_on_def
       
  1647   by fastforce
       
  1648 
       
  1649 lemma contour_integrable_div:
       
  1650     "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x / c) contour_integrable_on g"
       
  1651   using has_contour_integral_div contour_integrable_on_def
       
  1652   by fastforce
       
  1653 
       
  1654 lemma contour_integrable_setsum:
       
  1655     "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
       
  1656      \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) s) contour_integrable_on p"
       
  1657    unfolding contour_integrable_on_def
       
  1658    by (metis has_contour_integral_setsum)
       
  1659 
       
  1660 
       
  1661 subsection\<open>Reversing a path integral\<close>
       
  1662 
       
  1663 lemma has_contour_integral_reverse_linepath:
       
  1664     "(f has_contour_integral i) (linepath a b)
       
  1665      \<Longrightarrow> (f has_contour_integral (-i)) (linepath b a)"
       
  1666   using has_contour_integral_reversepath valid_path_linepath by fastforce
       
  1667 
       
  1668 lemma contour_integral_reverse_linepath:
       
  1669     "continuous_on (closed_segment a b) f
       
  1670      \<Longrightarrow> contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)"
       
  1671 apply (rule contour_integral_unique)
       
  1672 apply (rule has_contour_integral_reverse_linepath)
       
  1673 by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral)
       
  1674 
       
  1675 
       
  1676 (* Splitting a path integral in a flat way.*)
       
  1677 
       
  1678 lemma has_contour_integral_split:
       
  1679   assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)"
       
  1680       and k: "0 \<le> k" "k \<le> 1"
       
  1681       and c: "c - a = k *\<^sub>R (b - a)"
       
  1682     shows "(f has_contour_integral (i + j)) (linepath a b)"
       
  1683 proof (cases "k = 0 \<or> k = 1")
       
  1684   case True
       
  1685   then show ?thesis
       
  1686     using assms
       
  1687     apply auto
       
  1688     apply (metis add.left_neutral has_contour_integral_trivial has_contour_integral_unique)
       
  1689     apply (metis add.right_neutral has_contour_integral_trivial has_contour_integral_unique)
       
  1690     done
       
  1691 next
       
  1692   case False
       
  1693   then have k: "0 < k" "k < 1" "complex_of_real k \<noteq> 1"
       
  1694     using assms apply auto
       
  1695     using of_real_eq_iff by fastforce
       
  1696   have c': "c = k *\<^sub>R (b - a) + a"
       
  1697     by (metis diff_add_cancel c)
       
  1698   have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)"
       
  1699     by (simp add: algebra_simps c')
       
  1700   { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}"
       
  1701     have **: "\<And>x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b"
       
  1702       using False
       
  1703       apply (simp add: c' algebra_simps)
       
  1704       apply (simp add: real_vector.scale_left_distrib [symmetric] divide_simps)
       
  1705       done
       
  1706     have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}"
       
  1707       using * k
       
  1708       apply -
       
  1709       apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse k" "0", simplified])
       
  1710       apply (simp_all add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c)
       
  1711       apply (drule Integration.has_integral_cmul [where c = "inverse k"])
       
  1712       apply (simp add: Integration.has_integral_cmul)
       
  1713       done
       
  1714   } note fi = this
       
  1715   { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}"
       
  1716     have **: "\<And>x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)"
       
  1717       using k
       
  1718       apply (simp add: c' field_simps)
       
  1719       apply (simp add: scaleR_conv_of_real divide_simps)
       
  1720       apply (simp add: field_simps)
       
  1721       done
       
  1722     have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}"
       
  1723       using * k
       
  1724       apply -
       
  1725       apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse(1 - k)" "-(k/(1 - k))", simplified])
       
  1726       apply (simp_all add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc)
       
  1727       apply (drule Integration.has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"])
       
  1728       apply (simp add: Integration.has_integral_cmul)
       
  1729       done
       
  1730   } note fj = this
       
  1731   show ?thesis
       
  1732     using f k
       
  1733     apply (simp add: has_contour_integral_linepath)
       
  1734     apply (simp add: linepath_def)
       
  1735     apply (rule has_integral_combine [OF _ _ fi fj], simp_all)
       
  1736     done
       
  1737 qed
       
  1738 
       
  1739 lemma continuous_on_closed_segment_transform:
       
  1740   assumes f: "continuous_on (closed_segment a b) f"
       
  1741       and k: "0 \<le> k" "k \<le> 1"
       
  1742       and c: "c - a = k *\<^sub>R (b - a)"
       
  1743     shows "continuous_on (closed_segment a c) f"
       
  1744 proof -
       
  1745   have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b"
       
  1746     using c by (simp add: algebra_simps)
       
  1747   show "continuous_on (closed_segment a c) f"
       
  1748     apply (rule continuous_on_subset [OF f])
       
  1749     apply (simp add: segment_convex_hull)
       
  1750     apply (rule convex_hull_subset)
       
  1751     using assms
       
  1752     apply (auto simp: hull_inc c' Convex.convexD_alt)
       
  1753     done
       
  1754 qed
       
  1755 
       
  1756 lemma contour_integral_split:
       
  1757   assumes f: "continuous_on (closed_segment a b) f"
       
  1758       and k: "0 \<le> k" "k \<le> 1"
       
  1759       and c: "c - a = k *\<^sub>R (b - a)"
       
  1760     shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f"
       
  1761 proof -
       
  1762   have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b"
       
  1763     using c by (simp add: algebra_simps)
       
  1764   have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f"
       
  1765     apply (rule_tac [!] continuous_on_subset [OF f])
       
  1766     apply (simp_all add: segment_convex_hull)
       
  1767     apply (rule_tac [!] convex_hull_subset)
       
  1768     using assms
       
  1769     apply (auto simp: hull_inc c' Convex.convexD_alt)
       
  1770     done
       
  1771   show ?thesis
       
  1772     apply (rule contour_integral_unique)
       
  1773     apply (rule has_contour_integral_split [OF has_contour_integral_integral has_contour_integral_integral k c])
       
  1774     apply (rule contour_integrable_continuous_linepath *)+
       
  1775     done
       
  1776 qed
       
  1777 
       
  1778 lemma contour_integral_split_linepath:
       
  1779   assumes f: "continuous_on (closed_segment a b) f"
       
  1780       and c: "c \<in> closed_segment a b"
       
  1781     shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f"
       
  1782   using c
       
  1783   by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f])
       
  1784 
       
  1785 (* The special case of midpoints used in the main quadrisection.*)
       
  1786 
       
  1787 lemma has_contour_integral_midpoint:
       
  1788   assumes "(f has_contour_integral i) (linepath a (midpoint a b))"
       
  1789           "(f has_contour_integral j) (linepath (midpoint a b) b)"
       
  1790     shows "(f has_contour_integral (i + j)) (linepath a b)"
       
  1791   apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"])
       
  1792   using assms
       
  1793   apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real)
       
  1794   done
       
  1795 
       
  1796 lemma contour_integral_midpoint:
       
  1797    "continuous_on (closed_segment a b) f
       
  1798     \<Longrightarrow> contour_integral (linepath a b) f =
       
  1799         contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f"
       
  1800   apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"])
       
  1801   apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real)
       
  1802   done
       
  1803 
       
  1804 
       
  1805 text\<open>A couple of special case lemmas that are useful below\<close>
       
  1806 
       
  1807 lemma triangle_linear_has_chain_integral:
       
  1808     "((\<lambda>x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
       
  1809   apply (rule Cauchy_theorem_primitive [of UNIV "\<lambda>x. m/2 * x^2 + d*x"])
       
  1810   apply (auto intro!: derivative_eq_intros)
       
  1811   done
       
  1812 
       
  1813 lemma has_chain_integral_chain_integral3:
       
  1814      "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d)
       
  1815       \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i"
       
  1816   apply (subst contour_integral_unique [symmetric], assumption)
       
  1817   apply (drule has_contour_integral_integrable)
       
  1818   apply (simp add: valid_path_join)
       
  1819   done
       
  1820 
       
  1821 lemma has_chain_integral_chain_integral4:
       
  1822      "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e)
       
  1823       \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i"
       
  1824   apply (subst contour_integral_unique [symmetric], assumption)
       
  1825   apply (drule has_contour_integral_integrable)
       
  1826   apply (simp add: valid_path_join)
       
  1827   done
       
  1828 
       
  1829 subsection\<open>Reversing the order in a double path integral\<close>
       
  1830 
       
  1831 text\<open>The condition is stronger than needed but it's often true in typical situations\<close>
       
  1832 
       
  1833 lemma fst_im_cbox [simp]: "cbox c d \<noteq> {} \<Longrightarrow> (fst ` cbox (a,c) (b,d)) = cbox a b"
       
  1834   by (auto simp: cbox_Pair_eq)
       
  1835 
       
  1836 lemma snd_im_cbox [simp]: "cbox a b \<noteq> {} \<Longrightarrow> (snd ` cbox (a,c) (b,d)) = cbox c d"
       
  1837   by (auto simp: cbox_Pair_eq)
       
  1838 
       
  1839 lemma contour_integral_swap:
       
  1840   assumes fcon:  "continuous_on (path_image g \<times> path_image h) (\<lambda>(y1,y2). f y1 y2)"
       
  1841       and vp:    "valid_path g" "valid_path h"
       
  1842       and gvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative g (at t))"
       
  1843       and hvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative h (at t))"
       
  1844   shows "contour_integral g (\<lambda>w. contour_integral h (f w)) =
       
  1845          contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))"
       
  1846 proof -
       
  1847   have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h"
       
  1848     using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def)
       
  1849   have fgh1: "\<And>x. (\<lambda>t. f (g x) (h t)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g x, h t))"
       
  1850     by (rule ext) simp
       
  1851   have fgh2: "\<And>x. (\<lambda>t. f (g t) (h x)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g t, h x))"
       
  1852     by (rule ext) simp
       
  1853   have fcon_im1: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> continuous_on ((\<lambda>t. (g x, h t)) ` {0..1}) (\<lambda>(x, y). f x y)"
       
  1854     by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def)
       
  1855   have fcon_im2: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> continuous_on ((\<lambda>t. (g t, h x)) ` {0..1}) (\<lambda>(x, y). f x y)"
       
  1856     by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def)
       
  1857   have vdg: "\<And>y. y \<in> {0..1} \<Longrightarrow> (\<lambda>x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}"
       
  1858     apply (rule integrable_continuous_real)
       
  1859     apply (rule continuous_on_mult [OF _ gvcon])
       
  1860     apply (subst fgh2)
       
  1861     apply (rule fcon_im2 gcon continuous_intros | simp)+
       
  1862     done
       
  1863   have "(\<lambda>z. vector_derivative g (at (fst z))) = (\<lambda>x. vector_derivative g (at x)) o fst"
       
  1864     by auto
       
  1865   then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\<lambda>x. vector_derivative g (at (fst x)))"
       
  1866     apply (rule ssubst)
       
  1867     apply (rule continuous_intros | simp add: gvcon)+
       
  1868     done
       
  1869   have "(\<lambda>z. vector_derivative h (at (snd z))) = (\<lambda>x. vector_derivative h (at x)) o snd"
       
  1870     by auto
       
  1871   then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\<lambda>x. vector_derivative h (at (snd x)))"
       
  1872     apply (rule ssubst)
       
  1873     apply (rule continuous_intros | simp add: hvcon)+
       
  1874     done
       
  1875   have "(\<lambda>x. f (g (fst x)) (h (snd x))) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>w. ((g o fst) w, (h o snd) w))"
       
  1876     by auto
       
  1877   then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\<lambda>x. f (g (fst x)) (h (snd x)))"
       
  1878     apply (rule ssubst)
       
  1879     apply (rule gcon hcon continuous_intros | simp)+
       
  1880     apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon])
       
  1881     done
       
  1882   have "integral {0..1} (\<lambda>x. contour_integral h (f (g x)) * vector_derivative g (at x)) =
       
  1883         integral {0..1} (\<lambda>x. contour_integral h (\<lambda>y. f (g x) y * vector_derivative g (at x)))"
       
  1884     apply (rule integral_cong [OF contour_integral_rmul [symmetric]])
       
  1885     apply (clarsimp simp: contour_integrable_on)
       
  1886     apply (rule integrable_continuous_real)
       
  1887     apply (rule continuous_on_mult [OF _ hvcon])
       
  1888     apply (subst fgh1)
       
  1889     apply (rule fcon_im1 hcon continuous_intros | simp)+
       
  1890     done
       
  1891   also have "... = integral {0..1}
       
  1892                      (\<lambda>y. contour_integral g (\<lambda>x. f x (h y) * vector_derivative h (at y)))"
       
  1893     apply (simp only: contour_integral_integral)
       
  1894     apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified])
       
  1895      apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+
       
  1896     unfolding integral_mult_left [symmetric]
       
  1897     apply (simp only: mult_ac)
       
  1898     done
       
  1899   also have "... = contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))"
       
  1900     apply (simp add: contour_integral_integral)
       
  1901     apply (rule integral_cong)
       
  1902     unfolding integral_mult_left [symmetric]
       
  1903     apply (simp add: algebra_simps)
       
  1904     done
       
  1905   finally show ?thesis
       
  1906     by (simp add: contour_integral_integral)
       
  1907 qed
       
  1908 
       
  1909 
       
  1910 subsection\<open>The key quadrisection step\<close>
       
  1911 
       
  1912 lemma norm_sum_half:
       
  1913   assumes "norm(a + b) >= e"
       
  1914     shows "norm a >= e/2 \<or> norm b >= e/2"
       
  1915 proof -
       
  1916   have "e \<le> norm (- a - b)"
       
  1917     by (simp add: add.commute assms norm_minus_commute)
       
  1918   thus ?thesis
       
  1919     using norm_triangle_ineq4 order_trans by fastforce
       
  1920 qed
       
  1921 
       
  1922 lemma norm_sum_lemma:
       
  1923   assumes "e \<le> norm (a + b + c + d)"
       
  1924     shows "e / 4 \<le> norm a \<or> e / 4 \<le> norm b \<or> e / 4 \<le> norm c \<or> e / 4 \<le> norm d"
       
  1925 proof -
       
  1926   have "e \<le> norm ((a + b) + (c + d))" using assms
       
  1927     by (simp add: algebra_simps)
       
  1928   then show ?thesis
       
  1929     by (auto dest!: norm_sum_half)
       
  1930 qed
       
  1931 
       
  1932 lemma Cauchy_theorem_quadrisection:
       
  1933   assumes f: "continuous_on (convex hull {a,b,c}) f"
       
  1934       and dist: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
       
  1935       and e: "e * K^2 \<le>
       
  1936               norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)"
       
  1937   shows "\<exists>a' b' c'.
       
  1938            a' \<in> convex hull {a,b,c} \<and> b' \<in> convex hull {a,b,c} \<and> c' \<in> convex hull {a,b,c} \<and>
       
  1939            dist a' b' \<le> K/2  \<and>  dist b' c' \<le> K/2  \<and>  dist c' a' \<le> K/2  \<and>
       
  1940            e * (K/2)^2 \<le> norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)"
       
  1941 proof -
       
  1942   note divide_le_eq_numeral1 [simp del]
       
  1943   define a' where "a' = midpoint b c"
       
  1944   define b' where "b' = midpoint c a"
       
  1945   define c' where "c' = midpoint a b"
       
  1946   have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
       
  1947     using f continuous_on_subset segments_subset_convex_hull by metis+
       
  1948   have fcont': "continuous_on (closed_segment c' b') f"
       
  1949                "continuous_on (closed_segment a' c') f"
       
  1950                "continuous_on (closed_segment b' a') f"
       
  1951     unfolding a'_def b'_def c'_def
       
  1952     apply (rule continuous_on_subset [OF f],
       
  1953            metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+
       
  1954     done
       
  1955   let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
       
  1956   have *: "?pathint a b + ?pathint b c + ?pathint c a =
       
  1957           (?pathint a c' + ?pathint c' b' + ?pathint b' a) +
       
  1958           (?pathint a' c' + ?pathint c' b + ?pathint b a') +
       
  1959           (?pathint a' c + ?pathint c b' + ?pathint b' a') +
       
  1960           (?pathint a' b' + ?pathint b' c' + ?pathint c' a')"
       
  1961     apply (simp add: fcont' contour_integral_reverse_linepath)
       
  1962     apply (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc)
       
  1963     done
       
  1964   have [simp]: "\<And>x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2"
       
  1965     by (metis left_diff_distrib mult.commute norm_mult_numeral1)
       
  1966   have [simp]: "\<And>x y. cmod (x - y) = cmod (y - x)"
       
  1967     by (simp add: norm_minus_commute)
       
  1968   consider "e * K\<^sup>2 / 4 \<le> cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" |
       
  1969            "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" |
       
  1970            "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" |
       
  1971            "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')"
       
  1972     using assms
       
  1973     apply (simp only: *)
       
  1974     apply (blast intro: that dest!: norm_sum_lemma)
       
  1975     done
       
  1976   then show ?thesis
       
  1977   proof cases
       
  1978     case 1 then show ?thesis
       
  1979       apply (rule_tac x=a in exI)
       
  1980       apply (rule exI [where x=c'])
       
  1981       apply (rule exI [where x=b'])
       
  1982       using assms
       
  1983       apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
       
  1984       apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
       
  1985       done
       
  1986   next
       
  1987     case 2 then show ?thesis
       
  1988       apply (rule_tac x=a' in exI)
       
  1989       apply (rule exI [where x=c'])
       
  1990       apply (rule exI [where x=b])
       
  1991       using assms
       
  1992       apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
       
  1993       apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
       
  1994       done
       
  1995   next
       
  1996     case 3 then show ?thesis
       
  1997       apply (rule_tac x=a' in exI)
       
  1998       apply (rule exI [where x=c])
       
  1999       apply (rule exI [where x=b'])
       
  2000       using assms
       
  2001       apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
       
  2002       apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
       
  2003       done
       
  2004   next
       
  2005     case 4 then show ?thesis
       
  2006       apply (rule_tac x=a' in exI)
       
  2007       apply (rule exI [where x=b'])
       
  2008       apply (rule exI [where x=c'])
       
  2009       using assms
       
  2010       apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
       
  2011       apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
       
  2012       done
       
  2013   qed
       
  2014 qed
       
  2015 
       
  2016 subsection\<open>Cauchy's theorem for triangles\<close>
       
  2017 
       
  2018 lemma triangle_points_closer:
       
  2019   fixes a::complex
       
  2020   shows "\<lbrakk>x \<in> convex hull {a,b,c};  y \<in> convex hull {a,b,c}\<rbrakk>
       
  2021          \<Longrightarrow> norm(x - y) \<le> norm(a - b) \<or>
       
  2022              norm(x - y) \<le> norm(b - c) \<or>
       
  2023              norm(x - y) \<le> norm(c - a)"
       
  2024   using simplex_extremal_le [of "{a,b,c}"]
       
  2025   by (auto simp: norm_minus_commute)
       
  2026 
       
  2027 lemma holomorphic_point_small_triangle:
       
  2028   assumes x: "x \<in> s"
       
  2029       and f: "continuous_on s f"
       
  2030       and cd: "f field_differentiable (at x within s)"
       
  2031       and e: "0 < e"
       
  2032     shows "\<exists>k>0. \<forall>a b c. dist a b \<le> k \<and> dist b c \<le> k \<and> dist c a \<le> k \<and>
       
  2033               x \<in> convex hull {a,b,c} \<and> convex hull {a,b,c} \<subseteq> s
       
  2034               \<longrightarrow> norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f +
       
  2035                        contour_integral(linepath c a) f)
       
  2036                   \<le> e*(dist a b + dist b c + dist c a)^2"
       
  2037            (is "\<exists>k>0. \<forall>a b c. _ \<longrightarrow> ?normle a b c")
       
  2038 proof -
       
  2039   have le_of_3: "\<And>a x y z. \<lbrakk>0 \<le> x*y; 0 \<le> x*z; 0 \<le> y*z; a \<le> (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\<rbrakk>
       
  2040                      \<Longrightarrow> a \<le> e*(x + y + z)^2"
       
  2041     by (simp add: algebra_simps power2_eq_square)
       
  2042   have disj_le: "\<lbrakk>x \<le> a \<or> x \<le> b \<or> x \<le> c; 0 \<le> a; 0 \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> x \<le> a + b + c"
       
  2043              for x::real and a b c
       
  2044     by linarith
       
  2045   have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a"
       
  2046               if "convex hull {a, b, c} \<subseteq> s" for a b c
       
  2047     using segments_subset_convex_hull that
       
  2048     by (metis continuous_on_subset f contour_integrable_continuous_linepath)+
       
  2049   note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral]
       
  2050   { fix f' a b c d
       
  2051     assume d: "0 < d"
       
  2052        and f': "\<And>y. \<lbrakk>cmod (y - x) \<le> d; y \<in> s\<rbrakk> \<Longrightarrow> cmod (f y - f x - f' * (y - x)) \<le> e * cmod (y - x)"
       
  2053        and le: "cmod (a - b) \<le> d" "cmod (b - c) \<le> d" "cmod (c - a) \<le> d"
       
  2054        and xc: "x \<in> convex hull {a, b, c}"
       
  2055        and s: "convex hull {a, b, c} \<subseteq> s"
       
  2056     have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f =
       
  2057               contour_integral (linepath a b) (\<lambda>y. f y - f x - f'*(y - x)) +
       
  2058               contour_integral (linepath b c) (\<lambda>y. f y - f x - f'*(y - x)) +
       
  2059               contour_integral (linepath c a) (\<lambda>y. f y - f x - f'*(y - x))"
       
  2060       apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF s])
       
  2061       apply (simp add: field_simps)
       
  2062       done
       
  2063     { fix y
       
  2064       assume yc: "y \<in> convex hull {a,b,c}"
       
  2065       have "cmod (f y - f x - f' * (y - x)) \<le> e*norm(y - x)"
       
  2066         apply (rule f')
       
  2067         apply (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans)
       
  2068         using s yc by blast
       
  2069       also have "... \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))"
       
  2070         by (simp add: yc e xc disj_le [OF triangle_points_closer])
       
  2071       finally have "cmod (f y - f x - f' * (y - x)) \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" .
       
  2072     } note cm_le = this
       
  2073     have "?normle a b c"
       
  2074       apply (simp add: dist_norm pa)
       
  2075       apply (rule le_of_3)
       
  2076       using f' xc s e
       
  2077       apply simp_all
       
  2078       apply (intro norm_triangle_le add_mono path_bound)
       
  2079       apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc)
       
  2080       apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+
       
  2081       done
       
  2082   } note * = this
       
  2083   show ?thesis
       
  2084     using cd e
       
  2085     apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def)
       
  2086     apply (clarify dest!: spec mp)
       
  2087     using *
       
  2088     apply (simp add: dist_norm, blast)
       
  2089     done
       
  2090 qed
       
  2091 
       
  2092 
       
  2093 (* Hence the most basic theorem for a triangle.*)
       
  2094 locale Chain =
       
  2095   fixes x0 At Follows
       
  2096   assumes At0: "At x0 0"
       
  2097       and AtSuc: "\<And>x n. At x n \<Longrightarrow> \<exists>x'. At x' (Suc n) \<and> Follows x' x"
       
  2098 begin
       
  2099   primrec f where
       
  2100     "f 0 = x0"
       
  2101   | "f (Suc n) = (SOME x. At x (Suc n) \<and> Follows x (f n))"
       
  2102 
       
  2103   lemma At: "At (f n) n"
       
  2104   proof (induct n)
       
  2105     case 0 show ?case
       
  2106       by (simp add: At0)
       
  2107   next
       
  2108     case (Suc n) show ?case
       
  2109       by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex)
       
  2110   qed
       
  2111 
       
  2112   lemma Follows: "Follows (f(Suc n)) (f n)"
       
  2113     by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex)
       
  2114 
       
  2115   declare f.simps(2) [simp del]
       
  2116 end
       
  2117 
       
  2118 lemma Chain3:
       
  2119   assumes At0: "At x0 y0 z0 0"
       
  2120       and AtSuc: "\<And>x y z n. At x y z n \<Longrightarrow> \<exists>x' y' z'. At x' y' z' (Suc n) \<and> Follows x' y' z' x y z"
       
  2121   obtains f g h where
       
  2122     "f 0 = x0" "g 0 = y0" "h 0 = z0"
       
  2123                       "\<And>n. At (f n) (g n) (h n) n"
       
  2124                        "\<And>n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)"
       
  2125 proof -
       
  2126   interpret three: Chain "(x0,y0,z0)" "\<lambda>(x,y,z). At x y z" "\<lambda>(x',y',z'). \<lambda>(x,y,z). Follows x' y' z' x y z"
       
  2127     apply unfold_locales
       
  2128     using At0 AtSuc by auto
       
  2129   show ?thesis
       
  2130   apply (rule that [of "\<lambda>n. fst (three.f n)"  "\<lambda>n. fst (snd (three.f n))" "\<lambda>n. snd (snd (three.f n))"])
       
  2131   apply simp_all
       
  2132   using three.At three.Follows
       
  2133   apply (simp_all add: split_beta')
       
  2134   done
       
  2135 qed
       
  2136 
       
  2137 lemma Cauchy_theorem_triangle:
       
  2138   assumes "f holomorphic_on (convex hull {a,b,c})"
       
  2139     shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
       
  2140 proof -
       
  2141   have contf: "continuous_on (convex hull {a,b,c}) f"
       
  2142     by (metis assms holomorphic_on_imp_continuous_on)
       
  2143   let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
       
  2144   { fix y::complex
       
  2145     assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
       
  2146        and ynz: "y \<noteq> 0"
       
  2147     define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))"
       
  2148     define e where "e = norm y / K^2"
       
  2149     have K1: "K \<ge> 1"  by (simp add: K_def max.coboundedI1)
       
  2150     then have K: "K > 0" by linarith
       
  2151     have [iff]: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
       
  2152       by (simp_all add: K_def)
       
  2153     have e: "e > 0"
       
  2154       unfolding e_def using ynz K1 by simp
       
  2155     define At where "At x y z n \<longleftrightarrow>
       
  2156         convex hull {x,y,z} \<subseteq> convex hull {a,b,c} \<and>
       
  2157         dist x y \<le> K/2^n \<and> dist y z \<le> K/2^n \<and> dist z x \<le> K/2^n \<and>
       
  2158         norm(?pathint x y + ?pathint y z + ?pathint z x) \<ge> e*(K/2^n)^2"
       
  2159       for x y z n
       
  2160     have At0: "At a b c 0"
       
  2161       using fy
       
  2162       by (simp add: At_def e_def has_chain_integral_chain_integral3)
       
  2163     { fix x y z n
       
  2164       assume At: "At x y z n"
       
  2165       then have contf': "continuous_on (convex hull {x,y,z}) f"
       
  2166         using contf At_def continuous_on_subset by blast
       
  2167       have "\<exists>x' y' z'. At x' y' z' (Suc n) \<and> convex hull {x',y',z'} \<subseteq> convex hull {x,y,z}"
       
  2168         using At
       
  2169         apply (simp add: At_def)
       
  2170         using  Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e]
       
  2171         apply clarsimp
       
  2172         apply (rule_tac x="a'" in exI)
       
  2173         apply (rule_tac x="b'" in exI)
       
  2174         apply (rule_tac x="c'" in exI)
       
  2175         apply (simp add: algebra_simps)
       
  2176         apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE)
       
  2177         done
       
  2178     } note AtSuc = this
       
  2179     obtain fa fb fc
       
  2180       where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c"
       
  2181         and cosb: "\<And>n. convex hull {fa n, fb n, fc n} \<subseteq> convex hull {a,b,c}"
       
  2182         and dist: "\<And>n. dist (fa n) (fb n) \<le> K/2^n"
       
  2183                   "\<And>n. dist (fb n) (fc n) \<le> K/2^n"
       
  2184                   "\<And>n. dist (fc n) (fa n) \<le> K/2^n"
       
  2185         and no: "\<And>n. norm(?pathint (fa n) (fb n) +
       
  2186                            ?pathint (fb n) (fc n) +
       
  2187                            ?pathint (fc n) (fa n)) \<ge> e * (K/2^n)^2"
       
  2188         and conv_le: "\<And>n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \<subseteq> convex hull {fa n, fb n, fc n}"
       
  2189       apply (rule Chain3 [of At, OF At0 AtSuc])
       
  2190       apply (auto simp: At_def)
       
  2191       done
       
  2192     have "\<exists>x. \<forall>n. x \<in> convex hull {fa n, fb n, fc n}"
       
  2193       apply (rule bounded_closed_nest)
       
  2194       apply (simp_all add: compact_imp_closed finite_imp_compact_convex_hull finite_imp_bounded_convex_hull)
       
  2195       apply (rule allI)
       
  2196       apply (rule transitive_stepwise_le)
       
  2197       apply (auto simp: conv_le)
       
  2198       done
       
  2199     then obtain x where x: "\<And>n. x \<in> convex hull {fa n, fb n, fc n}" by auto
       
  2200     then have xin: "x \<in> convex hull {a,b,c}"
       
  2201       using assms f0 by blast
       
  2202     then have fx: "f field_differentiable at x within (convex hull {a,b,c})"
       
  2203       using assms holomorphic_on_def by blast
       
  2204     { fix k n
       
  2205       assume k: "0 < k"
       
  2206          and le:
       
  2207             "\<And>x' y' z'.
       
  2208                \<lbrakk>dist x' y' \<le> k; dist y' z' \<le> k; dist z' x' \<le> k;
       
  2209                 x \<in> convex hull {x',y',z'};
       
  2210                 convex hull {x',y',z'} \<subseteq> convex hull {a,b,c}\<rbrakk>
       
  2211                \<Longrightarrow>
       
  2212                cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10
       
  2213                      \<le> e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2"
       
  2214          and Kk: "K / k < 2 ^ n"
       
  2215       have "K / 2 ^ n < k" using Kk k
       
  2216         by (auto simp: field_simps)
       
  2217       then have DD: "dist (fa n) (fb n) \<le> k" "dist (fb n) (fc n) \<le> k" "dist (fc n) (fa n) \<le> k"
       
  2218         using dist [of n]  k
       
  2219         by linarith+
       
  2220       have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2
       
  2221                \<le> (3 * K / 2 ^ n)\<^sup>2"
       
  2222         using dist [of n] e K
       
  2223         by (simp add: abs_le_square_iff [symmetric])
       
  2224       have less10: "\<And>x y::real. 0 < x \<Longrightarrow> y \<le> 9*x \<Longrightarrow> y < x*10"
       
  2225         by linarith
       
  2226       have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \<le> e * (3 * K / 2 ^ n)\<^sup>2"
       
  2227         using ynz dle e mult_le_cancel_left_pos by blast
       
  2228       also have "... <
       
  2229           cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10"
       
  2230         using no [of n] e K
       
  2231         apply (simp add: e_def field_simps)
       
  2232         apply (simp only: zero_less_norm_iff [symmetric])
       
  2233         done
       
  2234       finally have False
       
  2235         using le [OF DD x cosb] by auto
       
  2236     } then
       
  2237     have ?thesis
       
  2238       using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e
       
  2239       apply clarsimp
       
  2240       apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]])
       
  2241       apply force+
       
  2242       done
       
  2243   }
       
  2244   moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
       
  2245     by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1)
       
  2246                    segments_subset_convex_hull(3) segments_subset_convex_hull(5))
       
  2247   ultimately show ?thesis
       
  2248     using has_contour_integral_integral by fastforce
       
  2249 qed
       
  2250 
       
  2251 
       
  2252 subsection\<open>Version needing function holomorphic in interior only\<close>
       
  2253 
       
  2254 lemma Cauchy_theorem_flat_lemma:
       
  2255   assumes f: "continuous_on (convex hull {a,b,c}) f"
       
  2256       and c: "c - a = k *\<^sub>R (b - a)"
       
  2257       and k: "0 \<le> k"
       
  2258     shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f +
       
  2259           contour_integral (linepath c a) f = 0"
       
  2260 proof -
       
  2261   have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
       
  2262     using f continuous_on_subset segments_subset_convex_hull by metis+
       
  2263   show ?thesis
       
  2264   proof (cases "k \<le> 1")
       
  2265     case True show ?thesis
       
  2266       by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc)
       
  2267   next
       
  2268     case False then show ?thesis
       
  2269       using fabc c
       
  2270       apply (subst contour_integral_split [of a c f "1/k" b, symmetric])
       
  2271       apply (metis closed_segment_commute fabc(3))
       
  2272       apply (auto simp: k contour_integral_reverse_linepath)
       
  2273       done
       
  2274   qed
       
  2275 qed
       
  2276 
       
  2277 lemma Cauchy_theorem_flat:
       
  2278   assumes f: "continuous_on (convex hull {a,b,c}) f"
       
  2279       and c: "c - a = k *\<^sub>R (b - a)"
       
  2280     shows "contour_integral (linepath a b) f +
       
  2281            contour_integral (linepath b c) f +
       
  2282            contour_integral (linepath c a) f = 0"
       
  2283 proof (cases "0 \<le> k")
       
  2284   case True with assms show ?thesis
       
  2285     by (blast intro: Cauchy_theorem_flat_lemma)
       
  2286 next
       
  2287   case False
       
  2288   have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
       
  2289     using f continuous_on_subset segments_subset_convex_hull by metis+
       
  2290   moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f +
       
  2291         contour_integral (linepath c b) f = 0"
       
  2292     apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"])
       
  2293     using False c
       
  2294     apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps)
       
  2295     done
       
  2296   ultimately show ?thesis
       
  2297     apply (auto simp: contour_integral_reverse_linepath)
       
  2298     using add_eq_0_iff by force
       
  2299 qed
       
  2300 
       
  2301 
       
  2302 lemma Cauchy_theorem_triangle_interior:
       
  2303   assumes contf: "continuous_on (convex hull {a,b,c}) f"
       
  2304       and holf:  "f holomorphic_on interior (convex hull {a,b,c})"
       
  2305      shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
       
  2306 proof -
       
  2307   have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
       
  2308     using contf continuous_on_subset segments_subset_convex_hull by metis+
       
  2309   have "bounded (f ` (convex hull {a,b,c}))"
       
  2310     by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf)
       
  2311   then obtain B where "0 < B" and Bnf: "\<And>x. x \<in> convex hull {a,b,c} \<Longrightarrow> norm (f x) \<le> B"
       
  2312      by (auto simp: dest!: bounded_pos [THEN iffD1])
       
  2313   have "bounded (convex hull {a,b,c})"
       
  2314     by (simp add: bounded_convex_hull)
       
  2315   then obtain C where C: "0 < C" and Cno: "\<And>y. y \<in> convex hull {a,b,c} \<Longrightarrow> norm y < C"
       
  2316     using bounded_pos_less by blast
       
  2317   then have diff_2C: "norm(x - y) \<le> 2*C"
       
  2318            if x: "x \<in> convex hull {a, b, c}" and y: "y \<in> convex hull {a, b, c}" for x y
       
  2319   proof -
       
  2320     have "cmod x \<le> C"
       
  2321       using x by (meson Cno not_le not_less_iff_gr_or_eq)
       
  2322     hence "cmod (x - y) \<le> C + C"
       
  2323       using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans)
       
  2324     thus "cmod (x - y) \<le> 2 * C"
       
  2325       by (metis mult_2)
       
  2326   qed
       
  2327   have contf': "continuous_on (convex hull {b,a,c}) f"
       
  2328     using contf by (simp add: insert_commute)
       
  2329   { fix y::complex
       
  2330     assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
       
  2331        and ynz: "y \<noteq> 0"
       
  2332     have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y"
       
  2333       by (rule has_chain_integral_chain_integral3 [OF fy])
       
  2334     have ?thesis
       
  2335     proof (cases "c=a \<or> a=b \<or> b=c")
       
  2336       case True then show ?thesis
       
  2337         using Cauchy_theorem_flat [OF contf, of 0]
       
  2338         using has_chain_integral_chain_integral3 [OF fy] ynz
       
  2339         by (force simp: fabc contour_integral_reverse_linepath)
       
  2340     next
       
  2341       case False
       
  2342       then have car3: "card {a, b, c} = Suc (DIM(complex))"
       
  2343         by auto
       
  2344       { assume "interior(convex hull {a,b,c}) = {}"
       
  2345         then have "collinear{a,b,c}"
       
  2346           using interior_convex_hull_eq_empty [OF car3]
       
  2347           by (simp add: collinear_3_eq_affine_dependent)
       
  2348         then have "False"
       
  2349           using False
       
  2350           apply (clarsimp simp add: collinear_3 collinear_lemma)
       
  2351           apply (drule Cauchy_theorem_flat [OF contf'])
       
  2352           using pi_eq_y ynz
       
  2353           apply (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath)
       
  2354           done
       
  2355       }
       
  2356       then obtain d where d: "d \<in> interior (convex hull {a, b, c})"
       
  2357         by blast
       
  2358       { fix d1
       
  2359         assume d1_pos: "0 < d1"
       
  2360            and d1: "\<And>x x'. \<lbrakk>x\<in>convex hull {a, b, c}; x'\<in>convex hull {a, b, c}; cmod (x' - x) < d1\<rbrakk>
       
  2361                            \<Longrightarrow> cmod (f x' - f x) < cmod y / (24 * C)"
       
  2362         define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))"
       
  2363         define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x
       
  2364         let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
       
  2365         have e: "0 < e" "e \<le> 1" "e \<le> d1 / (4 * C)" "e \<le> cmod y / 24 / C / B"
       
  2366           using d1_pos \<open>C>0\<close> \<open>B>0\<close> ynz by (simp_all add: e_def)
       
  2367         then have eCB: "24 * e * C * B \<le> cmod y"
       
  2368           using \<open>C>0\<close> \<open>B>0\<close>  by (simp add: field_simps)
       
  2369         have e_le_d1: "e * (4 * C) \<le> d1"
       
  2370           using e \<open>C>0\<close> by (simp add: field_simps)
       
  2371         have "shrink a \<in> interior(convex hull {a,b,c})"
       
  2372              "shrink b \<in> interior(convex hull {a,b,c})"
       
  2373              "shrink c \<in> interior(convex hull {a,b,c})"
       
  2374           using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def)
       
  2375         then have fhp0: "(f has_contour_integral 0)
       
  2376                 (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))"
       
  2377           by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal convex_interior)
       
  2378         then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0"
       
  2379           by (simp add: has_chain_integral_chain_integral3)
       
  2380         have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)"
       
  2381                       "f contour_integrable_on linepath (shrink b) (shrink c)"
       
  2382                       "f contour_integrable_on linepath (shrink c) (shrink a)"
       
  2383           using fhp0  by (auto simp: valid_path_join dest: has_contour_integral_integrable)
       
  2384         have cmod_shr: "\<And>x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)"
       
  2385           using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric])
       
  2386         have sh_eq: "\<And>a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)"
       
  2387           by (simp add: algebra_simps)
       
  2388         have "cmod y / (24 * C) \<le> cmod y / cmod (b - a) / 12"
       
  2389           using False \<open>C>0\<close> diff_2C [of b a] ynz
       
  2390           by (auto simp: divide_simps hull_inc)
       
  2391         have less_C: "\<lbrakk>u \<in> convex hull {a, b, c}; 0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> x * cmod u < C" for x u
       
  2392           apply (cases "x=0", simp add: \<open>0<C\<close>)
       
  2393           using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast
       
  2394         { fix u v
       
  2395           assume uv: "u \<in> convex hull {a, b, c}" "v \<in> convex hull {a, b, c}" "u\<noteq>v"
       
  2396              and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)"
       
  2397           have shr_uv: "shrink u \<in> interior(convex hull {a,b,c})"
       
  2398                        "shrink v \<in> interior(convex hull {a,b,c})"
       
  2399             using d e uv
       
  2400             by (auto simp: hull_inc mem_interior_convex_shrink shrink_def)
       
  2401           have cmod_fuv: "\<And>x. 0\<le>x \<Longrightarrow> x\<le>1 \<Longrightarrow> cmod (f (linepath (shrink u) (shrink v) x)) \<le> B"
       
  2402             using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD])
       
  2403           have By_uv: "B * (12 * (e * cmod (u - v))) \<le> cmod y"
       
  2404             apply (rule order_trans [OF _ eCB])
       
  2405             using e \<open>B>0\<close> diff_2C [of u v] uv
       
  2406             by (auto simp: field_simps)
       
  2407           { fix x::real   assume x: "0\<le>x" "x\<le>1"
       
  2408             have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)"
       
  2409               apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0])
       
  2410               using uv x d interior_subset
       
  2411               apply (auto simp: hull_inc intro!: less_C)
       
  2412               done
       
  2413             have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))"
       
  2414               by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real)
       
  2415             have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1"
       
  2416               using \<open>e>0\<close>
       
  2417               apply (simp add: ll norm_mult scaleR_diff_right)
       
  2418               apply (rule less_le_trans [OF _ e_le_d1])
       
  2419               using cmod_less_4C
       
  2420               apply (force intro: norm_triangle_lt)
       
  2421               done
       
  2422             have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)"
       
  2423               using x uv shr_uv cmod_less_dt
       
  2424               by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull)
       
  2425             also have "... \<le> cmod y / cmod (v - u) / 12"
       
  2426               using False uv \<open>C>0\<close> diff_2C [of v u] ynz
       
  2427               by (auto simp: divide_simps hull_inc)
       
  2428             finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \<le> cmod y / cmod (v - u) / 12"
       
  2429               by simp
       
  2430             then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \<le> cmod y"
       
  2431               using uv False by (auto simp: field_simps)
       
  2432             have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) +
       
  2433                   cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x))
       
  2434                   \<le> cmod y / 6"
       
  2435               apply (rule order_trans [of _ "B*((norm y / 24 / C / B)*2*C) + (2*C)*(norm y /24 / C)"])
       
  2436               apply (rule add_mono [OF mult_mono])
       
  2437               using By_uv e \<open>0 < B\<close> \<open>0 < C\<close> x ynz
       
  2438               apply (simp_all add: cmod_fuv cmod_shr cmod_12_le hull_inc)
       
  2439               apply (simp add: field_simps)
       
  2440               done
       
  2441           } note cmod_diff_le = this
       
  2442           have f_uv: "continuous_on (closed_segment u v) f"
       
  2443             by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull])
       
  2444           have **: "\<And>f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)"
       
  2445             by (simp add: algebra_simps)
       
  2446           have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \<le> norm y / 6"
       
  2447             apply (rule order_trans)
       
  2448             apply (rule has_integral_bound
       
  2449                     [of "B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)"
       
  2450                         "\<lambda>x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)"
       
  2451                         _ 0 1 ])
       
  2452             using ynz \<open>0 < B\<close> \<open>0 < C\<close>
       
  2453             apply (simp_all del: le_divide_eq_numeral1)
       
  2454             apply (simp add: has_integral_sub has_contour_integral_linepath [symmetric] has_contour_integral_integral
       
  2455                              fpi_uv f_uv contour_integrable_continuous_linepath, clarify)
       
  2456             apply (simp only: **)
       
  2457             apply (simp add: norm_triangle_le norm_mult cmod_diff_le del: le_divide_eq_numeral1)
       
  2458             done
       
  2459           } note * = this
       
  2460           have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \<le> norm y / 6"
       
  2461             using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
       
  2462           moreover
       
  2463           have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \<le> norm y / 6"
       
  2464             using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
       
  2465           moreover
       
  2466           have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \<le> norm y / 6"
       
  2467             using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
       
  2468           ultimately
       
  2469           have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) +
       
  2470                      (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a))
       
  2471                 \<le> norm y / 6 + norm y / 6 + norm y / 6"
       
  2472             by (metis norm_triangle_le add_mono)
       
  2473           also have "... = norm y / 2"
       
  2474             by simp
       
  2475           finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) -
       
  2476                           (?pathint a b + ?pathint b c + ?pathint c a))
       
  2477                 \<le> norm y / 2"
       
  2478             by (simp add: algebra_simps)
       
  2479           then
       
  2480           have "norm(?pathint a b + ?pathint b c + ?pathint c a) \<le> norm y / 2"
       
  2481             by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff)
       
  2482           then have "False"
       
  2483             using pi_eq_y ynz by auto
       
  2484         }
       
  2485         moreover have "uniformly_continuous_on (convex hull {a,b,c}) f"
       
  2486           by (simp add: contf compact_convex_hull compact_uniformly_continuous)
       
  2487         ultimately have "False"
       
  2488           unfolding uniformly_continuous_on_def
       
  2489           by (force simp: ynz \<open>0 < C\<close> dist_norm)
       
  2490         then show ?thesis ..
       
  2491       qed
       
  2492   }
       
  2493   moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
       
  2494     using fabc contour_integrable_continuous_linepath by auto
       
  2495   ultimately show ?thesis
       
  2496     using has_contour_integral_integral by fastforce
       
  2497 qed
       
  2498 
       
  2499 
       
  2500 subsection\<open>Version allowing finite number of exceptional points\<close>
       
  2501 
       
  2502 lemma Cauchy_theorem_triangle_cofinite:
       
  2503   assumes "continuous_on (convex hull {a,b,c}) f"
       
  2504       and "finite s"
       
  2505       and "(\<And>x. x \<in> interior(convex hull {a,b,c}) - s \<Longrightarrow> f field_differentiable (at x))"
       
  2506      shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
       
  2507 using assms
       
  2508 proof (induction "card s" arbitrary: a b c s rule: less_induct)
       
  2509   case (less s a b c)
       
  2510   show ?case
       
  2511   proof (cases "s={}")
       
  2512     case True with less show ?thesis
       
  2513       by (fastforce simp: holomorphic_on_def field_differentiable_at_within
       
  2514                     Cauchy_theorem_triangle_interior)
       
  2515   next
       
  2516     case False
       
  2517     then obtain d s' where d: "s = insert d s'" "d \<notin> s'"
       
  2518       by (meson Set.set_insert all_not_in_conv)
       
  2519     then show ?thesis
       
  2520     proof (cases "d \<in> convex hull {a,b,c}")
       
  2521       case False
       
  2522       show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
       
  2523         apply (rule less.hyps [of "s'"])
       
  2524         using False d \<open>finite s\<close> interior_subset
       
  2525         apply (auto intro!: less.prems)
       
  2526         done
       
  2527     next
       
  2528       case True
       
  2529       have *: "convex hull {a, b, d} \<subseteq> convex hull {a, b, c}"
       
  2530         by (meson True hull_subset insert_subset convex_hull_subset)
       
  2531       have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)"
       
  2532         apply (rule less.hyps [of "s'"])
       
  2533         using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
       
  2534         apply (auto intro!: less.prems continuous_on_subset [OF  _ *])
       
  2535         apply (metis * insert_absorb insert_subset interior_mono)
       
  2536         done
       
  2537       have *: "convex hull {b, c, d} \<subseteq> convex hull {a, b, c}"
       
  2538         by (meson True hull_subset insert_subset convex_hull_subset)
       
  2539       have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)"
       
  2540         apply (rule less.hyps [of "s'"])
       
  2541         using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
       
  2542         apply (auto intro!: less.prems continuous_on_subset [OF _ *])
       
  2543         apply (metis * insert_absorb insert_subset interior_mono)
       
  2544         done
       
  2545       have *: "convex hull {c, a, d} \<subseteq> convex hull {a, b, c}"
       
  2546         by (meson True hull_subset insert_subset convex_hull_subset)
       
  2547       have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)"
       
  2548         apply (rule less.hyps [of "s'"])
       
  2549         using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
       
  2550         apply (auto intro!: less.prems continuous_on_subset [OF _ *])
       
  2551         apply (metis * insert_absorb insert_subset interior_mono)
       
  2552         done
       
  2553       have "f contour_integrable_on linepath a b"
       
  2554         using less.prems
       
  2555         by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3))
       
  2556       moreover have "f contour_integrable_on linepath b c"
       
  2557         using less.prems
       
  2558         by (metis continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(3))
       
  2559       moreover have "f contour_integrable_on linepath c a"
       
  2560         using less.prems
       
  2561         by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3))
       
  2562       ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
       
  2563         by auto
       
  2564       { fix y::complex
       
  2565         assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
       
  2566            and ynz: "y \<noteq> 0"
       
  2567         have cont_ad: "continuous_on (closed_segment a d) f"
       
  2568           by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3))
       
  2569         have cont_bd: "continuous_on (closed_segment b d) f"
       
  2570           by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1))
       
  2571         have cont_cd: "continuous_on (closed_segment c d) f"
       
  2572           by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2))
       
  2573         have "contour_integral  (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))"
       
  2574                 "contour_integral  (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))"
       
  2575                 "contour_integral  (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)"
       
  2576             using has_chain_integral_chain_integral3 [OF abd]
       
  2577                   has_chain_integral_chain_integral3 [OF bcd]
       
  2578                   has_chain_integral_chain_integral3 [OF cad]
       
  2579             by (simp_all add: algebra_simps add_eq_0_iff)
       
  2580         then have ?thesis
       
  2581           using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce
       
  2582       }
       
  2583       then show ?thesis
       
  2584         using fpi contour_integrable_on_def by blast
       
  2585     qed
       
  2586   qed
       
  2587 qed
       
  2588 
       
  2589 
       
  2590 subsection\<open>Cauchy's theorem for an open starlike set\<close>
       
  2591 
       
  2592 lemma starlike_convex_subset:
       
  2593   assumes s: "a \<in> s" "closed_segment b c \<subseteq> s" and subs: "\<And>x. x \<in> s \<Longrightarrow> closed_segment a x \<subseteq> s"
       
  2594     shows "convex hull {a,b,c} \<subseteq> s"
       
  2595       using s
       
  2596       apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull)
       
  2597       apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE)
       
  2598       done
       
  2599 
       
  2600 lemma triangle_contour_integrals_starlike_primitive:
       
  2601   assumes contf: "continuous_on s f"
       
  2602       and s: "a \<in> s" "open s"
       
  2603       and x: "x \<in> s"
       
  2604       and subs: "\<And>y. y \<in> s \<Longrightarrow> closed_segment a y \<subseteq> s"
       
  2605       and zer: "\<And>b c. closed_segment b c \<subseteq> s
       
  2606                    \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f +
       
  2607                        contour_integral (linepath c a) f = 0"
       
  2608     shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)"
       
  2609 proof -
       
  2610   let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
       
  2611   { fix e y
       
  2612     assume e: "0 < e" and bxe: "ball x e \<subseteq> s" and close: "cmod (y - x) < e"
       
  2613     have y: "y \<in> s"
       
  2614       using bxe close  by (force simp: dist_norm norm_minus_commute)
       
  2615     have cont_ayf: "continuous_on (closed_segment a y) f"
       
  2616       using contf continuous_on_subset subs y by blast
       
  2617     have xys: "closed_segment x y \<subseteq> s"
       
  2618       apply (rule order_trans [OF _ bxe])
       
  2619       using close
       
  2620       by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound)
       
  2621     have "?pathint a y - ?pathint a x = ?pathint x y"
       
  2622       using zer [OF xys]  contour_integral_reverse_linepath [OF cont_ayf]  add_eq_0_iff by force
       
  2623   } note [simp] = this
       
  2624   { fix e::real
       
  2625     assume e: "0 < e"
       
  2626     have cont_atx: "continuous (at x) f"
       
  2627       using x s contf continuous_on_eq_continuous_at by blast
       
  2628     then obtain d1 where d1: "d1>0" and d1_less: "\<And>y. cmod (y - x) < d1 \<Longrightarrow> cmod (f y - f x) < e/2"
       
  2629       unfolding continuous_at Lim_at dist_norm  using e
       
  2630       by (drule_tac x="e/2" in spec) force
       
  2631     obtain d2 where d2: "d2>0" "ball x d2 \<subseteq> s" using  \<open>open s\<close> x
       
  2632       by (auto simp: open_contains_ball)
       
  2633     have dpos: "min d1 d2 > 0" using d1 d2 by simp
       
  2634     { fix y
       
  2635       assume yx: "y \<noteq> x" and close: "cmod (y - x) < min d1 d2"
       
  2636       have y: "y \<in> s"
       
  2637         using d2 close  by (force simp: dist_norm norm_minus_commute)
       
  2638       have fxy: "f contour_integrable_on linepath x y"
       
  2639         apply (rule contour_integrable_continuous_linepath)
       
  2640         apply (rule continuous_on_subset [OF contf])
       
  2641         using close d2
       
  2642         apply (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1))
       
  2643         done
       
  2644       then obtain i where i: "(f has_contour_integral i) (linepath x y)"
       
  2645         by (auto simp: contour_integrable_on_def)
       
  2646       then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)"
       
  2647         by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath])
       
  2648       then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)"
       
  2649         apply (rule has_contour_integral_bound_linepath [where B = "e/2"])
       
  2650         using e apply simp
       
  2651         apply (rule d1_less [THEN less_imp_le])
       
  2652         using close segment_bound
       
  2653         apply force
       
  2654         done
       
  2655       also have "... < e * cmod (y - x)"
       
  2656         by (simp add: e yx)
       
  2657       finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
       
  2658         using i yx  by (simp add: contour_integral_unique divide_less_eq)
       
  2659     }
       
  2660     then have "\<exists>d>0. \<forall>y. y \<noteq> x \<and> cmod (y-x) < d \<longrightarrow> cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
       
  2661       using dpos by blast
       
  2662   }
       
  2663   then have *: "(\<lambda>y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \<midarrow>x\<rightarrow> 0"
       
  2664     by (simp add: Lim_at dist_norm inverse_eq_divide)
       
  2665   show ?thesis
       
  2666     apply (simp add: has_field_derivative_def has_derivative_at bounded_linear_mult_right)
       
  2667     apply (rule Lim_transform [OF * Lim_eventually])
       
  2668     apply (simp add: inverse_eq_divide [symmetric] eventually_at)
       
  2669     using \<open>open s\<close> x
       
  2670     apply (force simp: dist_norm open_contains_ball)
       
  2671     done
       
  2672 qed
       
  2673 
       
  2674 (** Existence of a primitive.*)
       
  2675 
       
  2676 lemma holomorphic_starlike_primitive:
       
  2677   fixes f :: "complex \<Rightarrow> complex"
       
  2678   assumes contf: "continuous_on s f"
       
  2679       and s: "starlike s" and os: "open s"
       
  2680       and k: "finite k"
       
  2681       and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x"
       
  2682     shows "\<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x)"
       
  2683 proof -
       
  2684   obtain a where a: "a\<in>s" and a_cs: "\<And>x. x\<in>s \<Longrightarrow> closed_segment a x \<subseteq> s"
       
  2685     using s by (auto simp: starlike_def)
       
  2686   { fix x b c
       
  2687     assume "x \<in> s" "closed_segment b c \<subseteq> s"
       
  2688     then have abcs: "convex hull {a, b, c} \<subseteq> s"
       
  2689       by (simp add: a a_cs starlike_convex_subset)
       
  2690     then have *: "continuous_on (convex hull {a, b, c}) f"
       
  2691       by (simp add: continuous_on_subset [OF contf])
       
  2692     have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
       
  2693       apply (rule Cauchy_theorem_triangle_cofinite [OF _ k])
       
  2694       using abcs apply (simp add: continuous_on_subset [OF contf])
       
  2695       using * abcs interior_subset apply (auto intro: fcd)
       
  2696       done
       
  2697   } note 0 = this
       
  2698   show ?thesis
       
  2699     apply (intro exI ballI)
       
  2700     apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption)
       
  2701     apply (metis a_cs)
       
  2702     apply (metis has_chain_integral_chain_integral3 0)
       
  2703     done
       
  2704 qed
       
  2705 
       
  2706 lemma Cauchy_theorem_starlike:
       
  2707  "\<lbrakk>open s; starlike s; finite k; continuous_on s f;
       
  2708    \<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x;
       
  2709    valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk>
       
  2710    \<Longrightarrow> (f has_contour_integral 0)  g"
       
  2711   by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open)
       
  2712 
       
  2713 lemma Cauchy_theorem_starlike_simple:
       
  2714   "\<lbrakk>open s; starlike s; f holomorphic_on s; valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk>
       
  2715    \<Longrightarrow> (f has_contour_integral 0) g"
       
  2716 apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI])
       
  2717 apply (simp_all add: holomorphic_on_imp_continuous_on)
       
  2718 apply (metis at_within_open holomorphic_on_def)
       
  2719 done
       
  2720 
       
  2721 
       
  2722 subsection\<open>Cauchy's theorem for a convex set\<close>
       
  2723 
       
  2724 text\<open>For a convex set we can avoid assuming openness and boundary analyticity\<close>
       
  2725 
       
  2726 lemma triangle_contour_integrals_convex_primitive:
       
  2727   assumes contf: "continuous_on s f"
       
  2728       and s: "a \<in> s" "convex s"
       
  2729       and x: "x \<in> s"
       
  2730       and zer: "\<And>b c. \<lbrakk>b \<in> s; c \<in> s\<rbrakk>
       
  2731                    \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f +
       
  2732                        contour_integral (linepath c a) f = 0"
       
  2733     shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within s)"
       
  2734 proof -
       
  2735   let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
       
  2736   { fix y
       
  2737     assume y: "y \<in> s"
       
  2738     have cont_ayf: "continuous_on (closed_segment a y) f"
       
  2739       using s y  by (meson contf continuous_on_subset convex_contains_segment)
       
  2740     have xys: "closed_segment x y \<subseteq> s"  (*?*)
       
  2741       using convex_contains_segment s x y by auto
       
  2742     have "?pathint a y - ?pathint a x = ?pathint x y"
       
  2743       using zer [OF x y]  contour_integral_reverse_linepath [OF cont_ayf]  add_eq_0_iff by force
       
  2744   } note [simp] = this
       
  2745   { fix e::real
       
  2746     assume e: "0 < e"
       
  2747     have cont_atx: "continuous (at x within s) f"
       
  2748       using x s contf  by (simp add: continuous_on_eq_continuous_within)
       
  2749     then obtain d1 where d1: "d1>0" and d1_less: "\<And>y. \<lbrakk>y \<in> s; cmod (y - x) < d1\<rbrakk> \<Longrightarrow> cmod (f y - f x) < e/2"
       
  2750       unfolding continuous_within Lim_within dist_norm using e
       
  2751       by (drule_tac x="e/2" in spec) force
       
  2752     { fix y
       
  2753       assume yx: "y \<noteq> x" and close: "cmod (y - x) < d1" and y: "y \<in> s"
       
  2754       have fxy: "f contour_integrable_on linepath x y"
       
  2755         using convex_contains_segment s x y
       
  2756         by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
       
  2757       then obtain i where i: "(f has_contour_integral i) (linepath x y)"
       
  2758         by (auto simp: contour_integrable_on_def)
       
  2759       then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)"
       
  2760         by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath])
       
  2761       then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)"
       
  2762         apply (rule has_contour_integral_bound_linepath [where B = "e/2"])
       
  2763         using e apply simp
       
  2764         apply (rule d1_less [THEN less_imp_le])
       
  2765         using convex_contains_segment s(2) x y apply blast
       
  2766         using close segment_bound(1) apply fastforce
       
  2767         done
       
  2768       also have "... < e * cmod (y - x)"
       
  2769         by (simp add: e yx)
       
  2770       finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
       
  2771         using i yx  by (simp add: contour_integral_unique divide_less_eq)
       
  2772     }
       
  2773     then have "\<exists>d>0. \<forall>y\<in>s. y \<noteq> x \<and> cmod (y-x) < d \<longrightarrow> cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
       
  2774       using d1 by blast
       
  2775   }
       
  2776   then have *: "((\<lambda>y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \<longlongrightarrow> 0) (at x within s)"
       
  2777     by (simp add: Lim_within dist_norm inverse_eq_divide)
       
  2778   show ?thesis
       
  2779     apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right)
       
  2780     apply (rule Lim_transform [OF * Lim_eventually])
       
  2781     using linordered_field_no_ub
       
  2782     apply (force simp: inverse_eq_divide [symmetric] eventually_at)
       
  2783     done
       
  2784 qed
       
  2785 
       
  2786 lemma contour_integral_convex_primitive:
       
  2787   "\<lbrakk>convex s; continuous_on s f;
       
  2788     \<And>a b c. \<lbrakk>a \<in> s; b \<in> s; c \<in> s\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)\<rbrakk>
       
  2789          \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)"
       
  2790   apply (cases "s={}")
       
  2791   apply (simp_all add: ex_in_conv [symmetric])
       
  2792   apply (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3)
       
  2793   done
       
  2794 
       
  2795 lemma holomorphic_convex_primitive:
       
  2796   fixes f :: "complex \<Rightarrow> complex"
       
  2797   shows
       
  2798   "\<lbrakk>convex s; finite k; continuous_on s f;
       
  2799     \<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x\<rbrakk>
       
  2800    \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)"
       
  2801 apply (rule contour_integral_convex_primitive [OF _ _ Cauchy_theorem_triangle_cofinite])
       
  2802 prefer 3
       
  2803 apply (erule continuous_on_subset)
       
  2804 apply (simp add: subset_hull continuous_on_subset, assumption+)
       
  2805 by (metis Diff_iff convex_contains_segment insert_absorb insert_subset interior_mono segment_convex_hull subset_hull)
       
  2806 
       
  2807 lemma Cauchy_theorem_convex:
       
  2808     "\<lbrakk>continuous_on s f; convex s; finite k;
       
  2809       \<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x;
       
  2810      valid_path g; path_image g \<subseteq> s;
       
  2811      pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
       
  2812   by (metis holomorphic_convex_primitive Cauchy_theorem_primitive)
       
  2813 
       
  2814 lemma Cauchy_theorem_convex_simple:
       
  2815     "\<lbrakk>f holomorphic_on s; convex s;
       
  2816      valid_path g; path_image g \<subseteq> s;
       
  2817      pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
       
  2818   apply (rule Cauchy_theorem_convex)
       
  2819   apply (simp_all add: holomorphic_on_imp_continuous_on)
       
  2820   apply (rule finite.emptyI)
       
  2821   using at_within_interior holomorphic_on_def interior_subset by fastforce
       
  2822 
       
  2823 
       
  2824 text\<open>In particular for a disc\<close>
       
  2825 lemma Cauchy_theorem_disc:
       
  2826     "\<lbrakk>finite k; continuous_on (cball a e) f;
       
  2827       \<And>x. x \<in> ball a e - k \<Longrightarrow> f field_differentiable at x;
       
  2828      valid_path g; path_image g \<subseteq> cball a e;
       
  2829      pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
       
  2830   apply (rule Cauchy_theorem_convex)
       
  2831   apply (auto simp: convex_cball interior_cball)
       
  2832   done
       
  2833 
       
  2834 lemma Cauchy_theorem_disc_simple:
       
  2835     "\<lbrakk>f holomorphic_on (ball a e); valid_path g; path_image g \<subseteq> ball a e;
       
  2836      pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
       
  2837 by (simp add: Cauchy_theorem_convex_simple)
       
  2838 
       
  2839 
       
  2840 subsection\<open>Generalize integrability to local primitives\<close>
       
  2841 
       
  2842 lemma contour_integral_local_primitive_lemma:
       
  2843   fixes f :: "complex\<Rightarrow>complex"
       
  2844   shows
       
  2845     "\<lbrakk>g piecewise_differentiable_on {a..b};
       
  2846       \<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s);
       
  2847       \<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s\<rbrakk>
       
  2848      \<Longrightarrow> (\<lambda>x. f' (g x) * vector_derivative g (at x within {a..b}))
       
  2849             integrable_on {a..b}"
       
  2850   apply (cases "cbox a b = {}", force)
       
  2851   apply (simp add: integrable_on_def)
       
  2852   apply (rule exI)
       
  2853   apply (rule contour_integral_primitive_lemma, assumption+)
       
  2854   using atLeastAtMost_iff by blast
       
  2855 
       
  2856 lemma contour_integral_local_primitive_any:
       
  2857   fixes f :: "complex \<Rightarrow> complex"
       
  2858   assumes gpd: "g piecewise_differentiable_on {a..b}"
       
  2859       and dh: "\<And>x. x \<in> s
       
  2860                \<Longrightarrow> \<exists>d h. 0 < d \<and>
       
  2861                          (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
       
  2862       and gs: "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s"
       
  2863   shows "(\<lambda>x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}"
       
  2864 proof -
       
  2865   { fix x
       
  2866     assume x: "a \<le> x" "x \<le> b"
       
  2867     obtain d h where d: "0 < d"
       
  2868                and h: "(\<And>y. norm(y - g x) < d \<Longrightarrow> (h has_field_derivative f y) (at y within s))"
       
  2869       using x gs dh by (metis atLeastAtMost_iff)
       
  2870     have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast
       
  2871     then obtain e where e: "e>0" and lessd: "\<And>x'. x' \<in> {a..b} \<Longrightarrow> \<bar>x' - x\<bar> < e \<Longrightarrow> cmod (g x' - g x) < d"
       
  2872       using x d
       
  2873       apply (auto simp: dist_norm continuous_on_iff)
       
  2874       apply (drule_tac x=x in bspec)
       
  2875       using x apply simp
       
  2876       apply (drule_tac x=d in spec, auto)
       
  2877       done
       
  2878     have "\<exists>d>0. \<forall>u v. u \<le> x \<and> x \<le> v \<and> {u..v} \<subseteq> ball x d \<and> (u \<le> v \<longrightarrow> a \<le> u \<and> v \<le> b) \<longrightarrow>
       
  2879                           (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}"
       
  2880       apply (rule_tac x=e in exI)
       
  2881       using e
       
  2882       apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify)
       
  2883       apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma)
       
  2884         apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset)
       
  2885        apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force)
       
  2886       done
       
  2887   } then
       
  2888   show ?thesis
       
  2889     by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified])
       
  2890 qed
       
  2891 
       
  2892 lemma contour_integral_local_primitive:
       
  2893   fixes f :: "complex \<Rightarrow> complex"
       
  2894   assumes g: "valid_path g" "path_image g \<subseteq> s"
       
  2895       and dh: "\<And>x. x \<in> s
       
  2896                \<Longrightarrow> \<exists>d h. 0 < d \<and>
       
  2897                          (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
       
  2898   shows "f contour_integrable_on g"
       
  2899   using g
       
  2900   apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def
       
  2901             has_integral_localized_vector_derivative integrable_on_def [symmetric])
       
  2902   using contour_integral_local_primitive_any [OF _ dh]
       
  2903   by (meson image_subset_iff piecewise_C1_imp_differentiable)
       
  2904 
       
  2905 
       
  2906 text\<open>In particular if a function is holomorphic\<close>
       
  2907 
       
  2908 lemma contour_integrable_holomorphic:
       
  2909   assumes contf: "continuous_on s f"
       
  2910       and os: "open s"
       
  2911       and k: "finite k"
       
  2912       and g: "valid_path g" "path_image g \<subseteq> s"
       
  2913       and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x"
       
  2914     shows "f contour_integrable_on g"
       
  2915 proof -
       
  2916   { fix z
       
  2917     assume z: "z \<in> s"
       
  2918     obtain d where d: "d>0" "ball z d \<subseteq> s" using  \<open>open s\<close> z
       
  2919       by (auto simp: open_contains_ball)
       
  2920     then have contfb: "continuous_on (ball z d) f"
       
  2921       using contf continuous_on_subset by blast
       
  2922     obtain h where "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within ball z d)"
       
  2923       using holomorphic_convex_primitive [OF convex_ball k contfb fcd] d
       
  2924             interior_subset by force
       
  2925     then have "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within s)"
       
  2926       by (metis Topology_Euclidean_Space.open_ball at_within_open d(2) os subsetCE)
       
  2927     then have "\<exists>h. (\<forall>y. cmod (y - z) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
       
  2928       by (force simp: dist_norm norm_minus_commute)
       
  2929     then have "\<exists>d h. 0 < d \<and> (\<forall>y. cmod (y - z) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
       
  2930       using d by blast
       
  2931   }
       
  2932   then show ?thesis
       
  2933     by (rule contour_integral_local_primitive [OF g])
       
  2934 qed
       
  2935 
       
  2936 lemma contour_integrable_holomorphic_simple:
       
  2937   assumes fh: "f holomorphic_on s"
       
  2938       and os: "open s"
       
  2939       and g: "valid_path g" "path_image g \<subseteq> s"
       
  2940     shows "f contour_integrable_on g"
       
  2941   apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g])
       
  2942   apply (simp add: fh holomorphic_on_imp_continuous_on)
       
  2943   using fh  by (simp add: field_differentiable_def holomorphic_on_open os)
       
  2944 
       
  2945 lemma continuous_on_inversediff:
       
  2946   fixes z:: "'a::real_normed_field" shows "z \<notin> s \<Longrightarrow> continuous_on s (\<lambda>w. 1 / (w - z))"
       
  2947   by (rule continuous_intros | force)+
       
  2948 
       
  2949 corollary contour_integrable_inversediff:
       
  2950     "\<lbrakk>valid_path g; z \<notin> path_image g\<rbrakk> \<Longrightarrow> (\<lambda>w. 1 / (w-z)) contour_integrable_on g"
       
  2951 apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"])
       
  2952 apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros)
       
  2953 done
       
  2954 
       
  2955 text\<open>Key fact that path integral is the same for a "nearby" path. This is the
       
  2956  main lemma for the homotopy form of Cauchy's theorem and is also useful
       
  2957  if we want "without loss of generality" to assume some nice properties of a
       
  2958  path (e.g. smoothness). It can also be used to define the integrals of
       
  2959  analytic functions over arbitrary continuous paths. This is just done for
       
  2960  winding numbers now.
       
  2961 \<close>
       
  2962 
       
  2963 text\<open>A technical definition to avoid duplication of similar proofs,
       
  2964      for paths joined at the ends versus looping paths\<close>
       
  2965 definition linked_paths :: "bool \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
       
  2966   where "linked_paths atends g h ==
       
  2967         (if atends then pathstart h = pathstart g \<and> pathfinish h = pathfinish g
       
  2968                    else pathfinish g = pathstart g \<and> pathfinish h = pathstart h)"
       
  2969 
       
  2970 text\<open>This formulation covers two cases: @{term g} and @{term h} share their
       
  2971       start and end points; @{term g} and @{term h} both loop upon themselves.\<close>
       
  2972 lemma contour_integral_nearby:
       
  2973   assumes os: "open s" and p: "path p" "path_image p \<subseteq> s"
       
  2974     shows
       
  2975        "\<exists>d. 0 < d \<and>
       
  2976             (\<forall>g h. valid_path g \<and> valid_path h \<and>
       
  2977                   (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
       
  2978                   linked_paths atends g h
       
  2979                   \<longrightarrow> path_image g \<subseteq> s \<and> path_image h \<subseteq> s \<and>
       
  2980                       (\<forall>f. f holomorphic_on s \<longrightarrow> contour_integral h f = contour_integral g f))"
       
  2981 proof -
       
  2982   have "\<forall>z. \<exists>e. z \<in> path_image p \<longrightarrow> 0 < e \<and> ball z e \<subseteq> s"
       
  2983     using open_contains_ball os p(2) by blast
       
  2984   then obtain ee where ee: "\<And>z. z \<in> path_image p \<Longrightarrow> 0 < ee z \<and> ball z (ee z) \<subseteq> s"
       
  2985     by metis
       
  2986   define cover where "cover = (\<lambda>z. ball z (ee z/3)) ` (path_image p)"
       
  2987   have "compact (path_image p)"
       
  2988     by (metis p(1) compact_path_image)
       
  2989   moreover have "path_image p \<subseteq> (\<Union>c\<in>path_image p. ball c (ee c / 3))"
       
  2990     using ee by auto
       
  2991   ultimately have "\<exists>D \<subseteq> cover. finite D \<and> path_image p \<subseteq> \<Union>D"
       
  2992     by (simp add: compact_eq_heine_borel cover_def)
       
  2993   then obtain D where D: "D \<subseteq> cover" "finite D" "path_image p \<subseteq> \<Union>D"
       
  2994     by blast
       
  2995   then obtain k where k: "k \<subseteq> {0..1}" "finite k" and D_eq: "D = ((\<lambda>z. ball z (ee z / 3)) \<circ> p) ` k"
       
  2996     apply (simp add: cover_def path_image_def image_comp)
       
  2997     apply (blast dest!: finite_subset_image [OF \<open>finite D\<close>])
       
  2998     done
       
  2999   then have kne: "k \<noteq> {}"
       
  3000     using D by auto
       
  3001   have pi: "\<And>i. i \<in> k \<Longrightarrow> p i \<in> path_image p"
       
  3002     using k  by (auto simp: path_image_def)
       
  3003   then have eepi: "\<And>i. i \<in> k \<Longrightarrow> 0 < ee((p i))"
       
  3004     by (metis ee)
       
  3005   define e where "e = Min((ee o p) ` k)"
       
  3006   have fin_eep: "finite ((ee o p) ` k)"
       
  3007     using k  by blast
       
  3008   have enz: "0 < e"
       
  3009     using ee k  by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi)
       
  3010   have "uniformly_continuous_on {0..1} p"
       
  3011     using p  by (simp add: path_def compact_uniformly_continuous)
       
  3012   then obtain d::real where d: "d>0"
       
  3013           and de: "\<And>x x'. \<bar>x' - x\<bar> < d \<Longrightarrow> x\<in>{0..1} \<Longrightarrow> x'\<in>{0..1} \<Longrightarrow> cmod (p x' - p x) < e/3"
       
  3014     unfolding uniformly_continuous_on_def dist_norm real_norm_def
       
  3015     by (metis divide_pos_pos enz zero_less_numeral)
       
  3016   then obtain N::nat where N: "N>0" "inverse N < d"
       
  3017     using real_arch_inverse [of d]   by auto
       
  3018   { fix g h
       
  3019     assume g: "valid_path g" and gp: "\<forall>t\<in>{0..1}. cmod (g t - p t) < e / 3"
       
  3020        and h: "valid_path h" and hp: "\<forall>t\<in>{0..1}. cmod (h t - p t) < e / 3"
       
  3021        and joins: "linked_paths atends g h"
       
  3022     { fix t::real
       
  3023       assume t: "0 \<le> t" "t \<le> 1"
       
  3024       then obtain u where u: "u \<in> k" and ptu: "p t \<in> ball(p u) (ee(p u) / 3)"
       
  3025         using \<open>path_image p \<subseteq> \<Union>D\<close> D_eq by (force simp: path_image_def)
       
  3026       then have ele: "e \<le> ee (p u)" using fin_eep
       
  3027         by (simp add: e_def)
       
  3028       have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3"
       
  3029         using gp hp t by auto
       
  3030       with ele have "cmod (g t - p t) < ee (p u) / 3"
       
  3031                     "cmod (h t - p t) < ee (p u) / 3"
       
  3032         by linarith+
       
  3033       then have "g t \<in> ball(p u) (ee(p u))"  "h t \<in> ball(p u) (ee(p u))"
       
  3034         using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"]
       
  3035               norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u
       
  3036         by (force simp: dist_norm ball_def norm_minus_commute)+
       
  3037       then have "g t \<in> s" "h t \<in> s" using ee u k
       
  3038         by (auto simp: path_image_def ball_def)
       
  3039     }
       
  3040     then have ghs: "path_image g \<subseteq> s" "path_image h \<subseteq> s"
       
  3041       by (auto simp: path_image_def)
       
  3042     moreover
       
  3043     { fix f
       
  3044       assume fhols: "f holomorphic_on s"
       
  3045       then have fpa: "f contour_integrable_on g"  "f contour_integrable_on h"
       
  3046         using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple
       
  3047         by blast+
       
  3048       have contf: "continuous_on s f"
       
  3049         by (simp add: fhols holomorphic_on_imp_continuous_on)
       
  3050       { fix z
       
  3051         assume z: "z \<in> path_image p"
       
  3052         have "f holomorphic_on ball z (ee z)"
       
  3053           using fhols ee z holomorphic_on_subset by blast
       
  3054         then have "\<exists>ff. (\<forall>w \<in> ball z (ee z). (ff has_field_derivative f w) (at w))"
       
  3055           using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified]
       
  3056           by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball)
       
  3057       }
       
  3058       then obtain ff where ff:
       
  3059             "\<And>z w. \<lbrakk>z \<in> path_image p; w \<in> ball z (ee z)\<rbrakk> \<Longrightarrow> (ff z has_field_derivative f w) (at w)"
       
  3060         by metis
       
  3061       { fix n
       
  3062         assume n: "n \<le> N"
       
  3063         then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f =
       
  3064                    contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f"
       
  3065         proof (induct n)
       
  3066           case 0 show ?case by simp
       
  3067         next
       
  3068           case (Suc n)
       
  3069           obtain t where t: "t \<in> k" and "p (n/N) \<in> ball(p t) (ee(p t) / 3)"
       
  3070             using \<open>path_image p \<subseteq> \<Union>D\<close> [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems
       
  3071             by (force simp: path_image_def)
       
  3072           then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3"
       
  3073             by (simp add: dist_norm)
       
  3074           have e3le: "e/3 \<le> ee (p t) / 3"  using fin_eep t
       
  3075             by (simp add: e_def)
       
  3076           { fix x
       
  3077             assume x: "n/N \<le> x" "x \<le> (1 + n)/N"
       
  3078             then have nN01: "0 \<le> n/N" "(1 + n)/N \<le> 1"
       
  3079               using Suc.prems by auto
       
  3080             then have x01: "0 \<le> x" "x \<le> 1"
       
  3081               using x by linarith+
       
  3082             have "cmod (p t - p x)  < ee (p t) / 3 + e/3"
       
  3083               apply (rule norm_diff_triangle_less [OF ptu de])
       
  3084               using x N x01 Suc.prems
       
  3085               apply (auto simp: field_simps)
       
  3086               done
       
  3087             then have ptx: "cmod (p t - p x) < 2*ee (p t)/3"
       
  3088               using e3le eepi [OF t] by simp
       
  3089             have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 "
       
  3090               apply (rule norm_diff_triangle_less [OF ptx])
       
  3091               using gp x01 by (simp add: norm_minus_commute)
       
  3092             also have "... \<le> ee (p t)"
       
  3093               using e3le eepi [OF t] by simp
       
  3094             finally have gg: "cmod (p t - g x) < ee (p t)" .
       
  3095             have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 "
       
  3096               apply (rule norm_diff_triangle_less [OF ptx])
       
  3097               using hp x01 by (simp add: norm_minus_commute)
       
  3098             also have "... \<le> ee (p t)"
       
  3099               using e3le eepi [OF t] by simp
       
  3100             finally have "cmod (p t - g x) < ee (p t)"
       
  3101                          "cmod (p t - h x) < ee (p t)"
       
  3102               using gg by auto
       
  3103           } note ptgh_ee = this
       
  3104           have pi_hgn: "path_image (linepath (h (n/N)) (g (n/N))) \<subseteq> ball (p t) (ee (p t))"
       
  3105             using ptgh_ee [of "n/N"] Suc.prems
       
  3106             by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"])
       
  3107           then have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \<subseteq> s"
       
  3108             using \<open>N>0\<close> Suc.prems
       
  3109             apply (simp add: path_image_join field_simps closed_segment_commute)
       
  3110             apply (erule order_trans)
       
  3111             apply (simp add: ee pi t)
       
  3112             done
       
  3113           have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N)))
       
  3114                   \<subseteq> ball (p t) (ee (p t))"
       
  3115             using ptgh_ee [of "(1+n)/N"] Suc.prems
       
  3116             by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"])
       
  3117           then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \<subseteq> s"
       
  3118             using \<open>N>0\<close> Suc.prems ee pi t
       
  3119             by (auto simp: Path_Connected.path_image_join field_simps)
       
  3120           have pi_subset_ball:
       
  3121                 "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++
       
  3122                              subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N)))
       
  3123                  \<subseteq> ball (p t) (ee (p t))"
       
  3124             apply (intro subset_path_image_join pi_hgn pi_ghn')
       
  3125             using \<open>N>0\<close> Suc.prems
       
  3126             apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee)
       
  3127             done
       
  3128           have pi0: "(f has_contour_integral 0)
       
  3129                        (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++
       
  3130                         subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))"
       
  3131             apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"])
       
  3132             apply (metis ff open_ball at_within_open pi t)
       
  3133             apply (intro valid_path_join)
       
  3134             using Suc.prems pi_subset_ball apply (simp_all add: valid_path_subpath g h)
       
  3135             done
       
  3136           have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g"
       
  3137             using Suc.prems by (simp add: contour_integrable_subpath g fpa)
       
  3138           have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))"
       
  3139             using gh_n's
       
  3140             by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
       
  3141           have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))"
       
  3142             using gh_ns
       
  3143             by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
       
  3144           have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f +
       
  3145                      contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f +
       
  3146                      contour_integral (subpath ((Suc n) / N) (n/N) h) f +
       
  3147                      contour_integral (linepath (h (n/N)) (g (n/N))) f = 0"
       
  3148             using contour_integral_unique [OF pi0] Suc.prems
       
  3149             by (simp add: g h fpa valid_path_subpath contour_integrable_subpath
       
  3150                           fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc)
       
  3151           have *: "\<And>hn he hn' gn gd gn' hgn ghn gh0 ghn'.
       
  3152                     \<lbrakk>hn - gn = ghn - gh0;
       
  3153                      gd + ghn' + he + hgn = (0::complex);
       
  3154                      hn - he = hn'; gn + gd = gn'; hgn = -ghn\<rbrakk> \<Longrightarrow> hn' - gn' = ghn' - gh0"
       
  3155             by (auto simp: algebra_simps)
       
  3156           have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f =
       
  3157                 contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f"
       
  3158             unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"]
       
  3159             using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath)
       
  3160           also have "... = contour_integral (subpath 0 ((Suc n) / N) h) f"
       
  3161             using Suc.prems by (simp add: contour_integral_subpath_combine h fpa)
       
  3162           finally have pi0_eq:
       
  3163                "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f =
       
  3164                 contour_integral (subpath 0 ((Suc n) / N) h) f" .
       
  3165           show ?case
       
  3166             apply (rule * [OF Suc.hyps eq0 pi0_eq])
       
  3167             using Suc.prems
       
  3168             apply (simp_all add: g h fpa contour_integral_subpath_combine
       
  3169                      contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath
       
  3170                      continuous_on_subset [OF contf gh_ns])
       
  3171             done
       
  3172       qed
       
  3173       } note ind = this
       
  3174       have "contour_integral h f = contour_integral g f"
       
  3175         using ind [OF order_refl] N joins
       
  3176         by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm)
       
  3177     }
       
  3178     ultimately
       
  3179     have "path_image g \<subseteq> s \<and> path_image h \<subseteq> s \<and> (\<forall>f. f holomorphic_on s \<longrightarrow> contour_integral h f = contour_integral g f)"
       
  3180       by metis
       
  3181   } note * = this
       
  3182   show ?thesis
       
  3183     apply (rule_tac x="e/3" in exI)
       
  3184     apply (rule conjI)
       
  3185     using enz apply simp
       
  3186     apply (clarsimp simp only: ball_conj_distrib)
       
  3187     apply (rule *; assumption)
       
  3188     done
       
  3189 qed
       
  3190 
       
  3191 
       
  3192 lemma
       
  3193   assumes "open s" "path p" "path_image p \<subseteq> s"
       
  3194     shows contour_integral_nearby_ends:
       
  3195       "\<exists>d. 0 < d \<and>
       
  3196               (\<forall>g h. valid_path g \<and> valid_path h \<and>
       
  3197                     (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
       
  3198                     pathstart h = pathstart g \<and> pathfinish h = pathfinish g
       
  3199                     \<longrightarrow> path_image g \<subseteq> s \<and>
       
  3200                         path_image h \<subseteq> s \<and>
       
  3201                         (\<forall>f. f holomorphic_on s
       
  3202                             \<longrightarrow> contour_integral h f = contour_integral g f))"
       
  3203     and contour_integral_nearby_loops:
       
  3204       "\<exists>d. 0 < d \<and>
       
  3205               (\<forall>g h. valid_path g \<and> valid_path h \<and>
       
  3206                     (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
       
  3207                     pathfinish g = pathstart g \<and> pathfinish h = pathstart h
       
  3208                     \<longrightarrow> path_image g \<subseteq> s \<and>
       
  3209                         path_image h \<subseteq> s \<and>
       
  3210                         (\<forall>f. f holomorphic_on s
       
  3211                             \<longrightarrow> contour_integral h f = contour_integral g f))"
       
  3212   using contour_integral_nearby [OF assms, where atends=True]
       
  3213   using contour_integral_nearby [OF assms, where atends=False]
       
  3214   unfolding linked_paths_def by simp_all
       
  3215 
       
  3216 corollary differentiable_polynomial_function:
       
  3217   fixes p :: "real \<Rightarrow> 'a::euclidean_space"
       
  3218   shows "polynomial_function p \<Longrightarrow> p differentiable_on s"
       
  3219 by (meson has_vector_derivative_polynomial_function differentiable_at_imp_differentiable_on differentiable_def has_vector_derivative_def)
       
  3220 
       
  3221 lemma C1_differentiable_polynomial_function:
       
  3222   fixes p :: "real \<Rightarrow> 'a::euclidean_space"
       
  3223   shows "polynomial_function p \<Longrightarrow> p C1_differentiable_on s"
       
  3224   by (metis continuous_on_polymonial_function C1_differentiable_on_def  has_vector_derivative_polynomial_function)
       
  3225 
       
  3226 lemma valid_path_polynomial_function:
       
  3227   fixes p :: "real \<Rightarrow> 'a::euclidean_space"
       
  3228   shows "polynomial_function p \<Longrightarrow> valid_path p"
       
  3229 by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function)
       
  3230 
       
  3231 lemma valid_path_subpath_trivial [simp]:
       
  3232     fixes g :: "real \<Rightarrow> 'a::euclidean_space"
       
  3233     shows "z \<noteq> g x \<Longrightarrow> valid_path (subpath x x g)"
       
  3234   by (simp add: subpath_def valid_path_polynomial_function)
       
  3235 
       
  3236 lemma contour_integral_bound_exists:
       
  3237 assumes s: "open s"
       
  3238     and g: "valid_path g"
       
  3239     and pag: "path_image g \<subseteq> s"
       
  3240   shows "\<exists>L. 0 < L \<and>
       
  3241        (\<forall>f B. f holomorphic_on s \<and> (\<forall>z \<in> s. norm(f z) \<le> B)
       
  3242          \<longrightarrow> norm(contour_integral g f) \<le> L*B)"
       
  3243 proof -
       
  3244 have "path g" using g
       
  3245   by (simp add: valid_path_imp_path)
       
  3246 then obtain d::real and p
       
  3247   where d: "0 < d"
       
  3248     and p: "polynomial_function p" "path_image p \<subseteq> s"
       
  3249     and pi: "\<And>f. f holomorphic_on s \<Longrightarrow> contour_integral g f = contour_integral p f"
       
  3250   using contour_integral_nearby_ends [OF s \<open>path g\<close> pag]
       
  3251   apply clarify
       
  3252   apply (drule_tac x=g in spec)
       
  3253   apply (simp only: assms)
       
  3254   apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function)
       
  3255   done
       
  3256 then obtain p' where p': "polynomial_function p'"
       
  3257          "\<And>x. (p has_vector_derivative (p' x)) (at x)"
       
  3258   using has_vector_derivative_polynomial_function by force
       
  3259 then have "bounded(p' ` {0..1})"
       
  3260   using continuous_on_polymonial_function
       
  3261   by (force simp: intro!: compact_imp_bounded compact_continuous_image)
       
  3262 then obtain L where L: "L>0" and nop': "\<And>x. x \<in> {0..1} \<Longrightarrow> norm (p' x) \<le> L"
       
  3263   by (force simp: bounded_pos)
       
  3264 { fix f B
       
  3265   assume f: "f holomorphic_on s"
       
  3266      and B: "\<And>z. z\<in>s \<Longrightarrow> cmod (f z) \<le> B"
       
  3267   then have "f contour_integrable_on p \<and> valid_path p"
       
  3268     using p s
       
  3269     by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on)
       
  3270   moreover have "\<And>x. x \<in> {0..1} \<Longrightarrow> cmod (vector_derivative p (at x)) * cmod (f (p x)) \<le> L * B"
       
  3271     apply (rule mult_mono)
       
  3272     apply (subst Derivative.vector_derivative_at; force intro: p' nop')
       
  3273     using L B p
       
  3274     apply (auto simp: path_image_def image_subset_iff)
       
  3275     done
       
  3276   ultimately have "cmod (contour_integral g f) \<le> L * B"
       
  3277     apply (simp add: pi [OF f])
       
  3278     apply (simp add: contour_integral_integral)
       
  3279     apply (rule order_trans [OF integral_norm_bound_integral])
       
  3280     apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult)
       
  3281     done
       
  3282 } then
       
  3283 show ?thesis
       
  3284   by (force simp: L contour_integral_integral)
       
  3285 qed
       
  3286 
       
  3287 subsection\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close>
       
  3288 
       
  3289 text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close>
       
  3290 
       
  3291 lemma continuous_disconnected_range_constant:
       
  3292   assumes s: "connected s"
       
  3293       and conf: "continuous_on s f"
       
  3294       and fim: "f ` s \<subseteq> t"
       
  3295       and cct: "\<And>y. y \<in> t \<Longrightarrow> connected_component_set t y = {y}"
       
  3296     shows "\<exists>a. \<forall>x \<in> s. f x = a"
       
  3297 proof (cases "s = {}")
       
  3298   case True then show ?thesis by force
       
  3299 next
       
  3300   case False
       
  3301   { fix x assume "x \<in> s"
       
  3302     then have "f ` s \<subseteq> {f x}"
       
  3303     by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI s cct)
       
  3304   }
       
  3305   with False show ?thesis
       
  3306     by blast
       
  3307 qed
       
  3308 
       
  3309 lemma discrete_subset_disconnected:
       
  3310   fixes s :: "'a::topological_space set"
       
  3311   fixes t :: "'b::real_normed_vector set"
       
  3312   assumes conf: "continuous_on s f"
       
  3313       and no: "\<And>x. x \<in> s \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)"
       
  3314    shows "f ` s \<subseteq> {y. connected_component_set (f ` s) y = {y}}"
       
  3315 proof -
       
  3316   { fix x assume x: "x \<in> s"
       
  3317     then obtain e where "e>0" and ele: "\<And>y. \<lbrakk>y \<in> s; f y \<noteq> f x\<rbrakk> \<Longrightarrow> e \<le> norm (f y - f x)"
       
  3318       using conf no [OF x] by auto
       
  3319     then have e2: "0 \<le> e / 2"
       
  3320       by simp
       
  3321     have "f y = f x" if "y \<in> s" and ccs: "f y \<in> connected_component_set (f ` s) (f x)" for y
       
  3322       apply (rule ccontr)
       
  3323       using connected_closed [of "connected_component_set (f ` s) (f x)"] \<open>e>0\<close>
       
  3324       apply (simp add: del: ex_simps)
       
  3325       apply (drule spec [where x="cball (f x) (e / 2)"])
       
  3326       apply (drule spec [where x="- ball(f x) e"])
       
  3327       apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in)
       
  3328         apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less)
       
  3329        using centre_in_cball connected_component_refl_eq e2 x apply blast
       
  3330       using ccs
       
  3331       apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \<open>y \<in> s\<close>])
       
  3332       done
       
  3333     moreover have "connected_component_set (f ` s) (f x) \<subseteq> f ` s"
       
  3334       by (auto simp: connected_component_in)
       
  3335     ultimately have "connected_component_set (f ` s) (f x) = {f x}"
       
  3336       by (auto simp: x)
       
  3337   }
       
  3338   with assms show ?thesis
       
  3339     by blast
       
  3340 qed
       
  3341 
       
  3342 lemma finite_implies_discrete:
       
  3343   fixes s :: "'a::topological_space set"
       
  3344   assumes "finite (f ` s)"
       
  3345   shows "(\<forall>x \<in> s. \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x))"
       
  3346 proof -
       
  3347   have "\<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" if "x \<in> s" for x
       
  3348   proof (cases "f ` s - {f x} = {}")
       
  3349     case True
       
  3350     with zero_less_numeral show ?thesis
       
  3351       by (fastforce simp add: Set.image_subset_iff cong: conj_cong)
       
  3352   next
       
  3353     case False
       
  3354     then obtain z where z: "z \<in> s" "f z \<noteq> f x"
       
  3355       by blast
       
  3356     have finn: "finite {norm (z - f x) |z. z \<in> f ` s - {f x}}"
       
  3357       using assms by simp
       
  3358     then have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` s - {f x}}"
       
  3359       apply (rule finite_imp_less_Inf)
       
  3360       using z apply force+
       
  3361       done
       
  3362     show ?thesis
       
  3363       by (force intro!: * cInf_le_finite [OF finn])
       
  3364   qed
       
  3365   with assms show ?thesis
       
  3366     by blast
       
  3367 qed
       
  3368 
       
  3369 text\<open>This proof requires the existence of two separate values of the range type.\<close>
       
  3370 lemma finite_range_constant_imp_connected:
       
  3371   assumes "\<And>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
       
  3372               \<lbrakk>continuous_on s f; finite(f ` s)\<rbrakk> \<Longrightarrow> \<exists>a. \<forall>x \<in> s. f x = a"
       
  3373     shows "connected s"
       
  3374 proof -
       
  3375   { fix t u
       
  3376     assume clt: "closedin (subtopology euclidean s) t"
       
  3377        and clu: "closedin (subtopology euclidean s) u"
       
  3378        and tue: "t \<inter> u = {}" and tus: "t \<union> u = s"
       
  3379     have conif: "continuous_on s (\<lambda>x. if x \<in> t then 0 else 1)"
       
  3380       apply (subst tus [symmetric])
       
  3381       apply (rule continuous_on_cases_local)
       
  3382       using clt clu tue
       
  3383       apply (auto simp: tus continuous_on_const)
       
  3384       done
       
  3385     have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` s)"
       
  3386       by (rule finite_subset [of _ "{0,1}"]) auto
       
  3387     have "t = {} \<or> u = {}"
       
  3388       using assms [OF conif fi] tus [symmetric]
       
  3389       by (auto simp: Ball_def) (metis IntI empty_iff one_neq_zero tue)
       
  3390   }
       
  3391   then show ?thesis
       
  3392     by (simp add: connected_closedin_eq)
       
  3393 qed
       
  3394 
       
  3395 lemma continuous_disconnected_range_constant_eq:
       
  3396       "(connected s \<longleftrightarrow>
       
  3397            (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
       
  3398             \<forall>t. continuous_on s f \<and> f ` s \<subseteq> t \<and> (\<forall>y \<in> t. connected_component_set t y = {y})
       
  3399             \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis1)
       
  3400   and continuous_discrete_range_constant_eq:
       
  3401       "(connected s \<longleftrightarrow>
       
  3402          (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
       
  3403           continuous_on s f \<and>
       
  3404           (\<forall>x \<in> s. \<exists>e. 0 < e \<and> (\<forall>y. y \<in> s \<and> (f y \<noteq> f x) \<longrightarrow> e \<le> norm(f y - f x)))
       
  3405           \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis2)
       
  3406   and continuous_finite_range_constant_eq:
       
  3407       "(connected s \<longleftrightarrow>
       
  3408          (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
       
  3409           continuous_on s f \<and> finite (f ` s)
       
  3410           \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis3)
       
  3411 proof -
       
  3412   have *: "\<And>s t u v. \<lbrakk>s \<Longrightarrow> t; t \<Longrightarrow> u; u \<Longrightarrow> v; v \<Longrightarrow> s\<rbrakk>
       
  3413     \<Longrightarrow> (s \<longleftrightarrow> t) \<and> (s \<longleftrightarrow> u) \<and> (s \<longleftrightarrow> v)"
       
  3414     by blast
       
  3415   have "?thesis1 \<and> ?thesis2 \<and> ?thesis3"
       
  3416     apply (rule *)
       
  3417     using continuous_disconnected_range_constant apply metis
       
  3418     apply clarify
       
  3419     apply (frule discrete_subset_disconnected; blast)
       
  3420     apply (blast dest: finite_implies_discrete)
       
  3421     apply (blast intro!: finite_range_constant_imp_connected)
       
  3422     done
       
  3423   then show ?thesis1 ?thesis2 ?thesis3
       
  3424     by blast+
       
  3425 qed
       
  3426 
       
  3427 lemma continuous_discrete_range_constant:
       
  3428   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1"
       
  3429   assumes s: "connected s"
       
  3430       and "continuous_on s f"
       
  3431       and "\<And>x. x \<in> s \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)"
       
  3432     shows "\<exists>a. \<forall>x \<in> s. f x = a"
       
  3433   using continuous_discrete_range_constant_eq [THEN iffD1, OF s] assms
       
  3434   by blast
       
  3435 
       
  3436 lemma continuous_finite_range_constant:
       
  3437   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1"
       
  3438   assumes "connected s"
       
  3439       and "continuous_on s f"
       
  3440       and "finite (f ` s)"
       
  3441     shows "\<exists>a. \<forall>x \<in> s. f x = a"
       
  3442   using assms continuous_finite_range_constant_eq
       
  3443   by blast
       
  3444 
       
  3445 
       
  3446 text\<open>We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\<close>
       
  3447 
       
  3448 subsection\<open>Winding Numbers\<close>
       
  3449 
       
  3450 definition winding_number:: "[real \<Rightarrow> complex, complex] \<Rightarrow> complex" where
       
  3451   "winding_number \<gamma> z \<equiv>
       
  3452     @n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
       
  3453                     pathstart p = pathstart \<gamma> \<and>
       
  3454                     pathfinish p = pathfinish \<gamma> \<and>
       
  3455                     (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
       
  3456                     contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
       
  3457 
       
  3458 lemma winding_number:
       
  3459   assumes "path \<gamma>" "z \<notin> path_image \<gamma>" "0 < e"
       
  3460     shows "\<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
       
  3461                pathstart p = pathstart \<gamma> \<and>
       
  3462                pathfinish p = pathfinish \<gamma> \<and>
       
  3463                (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
       
  3464                contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * winding_number \<gamma> z"
       
  3465 proof -
       
  3466   have "path_image \<gamma> \<subseteq> UNIV - {z}"
       
  3467     using assms by blast
       
  3468   then obtain d
       
  3469     where d: "d>0"
       
  3470       and pi_eq: "\<And>h1 h2. valid_path h1 \<and> valid_path h2 \<and>
       
  3471                     (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d) \<and>
       
  3472                     pathstart h2 = pathstart h1 \<and> pathfinish h2 = pathfinish h1 \<longrightarrow>
       
  3473                       path_image h1 \<subseteq> UNIV - {z} \<and> path_image h2 \<subseteq> UNIV - {z} \<and>
       
  3474                       (\<forall>f. f holomorphic_on UNIV - {z} \<longrightarrow> contour_integral h2 f = contour_integral h1 f)"
       
  3475     using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms by (auto simp: open_delete)
       
  3476   then obtain h where h: "polynomial_function h \<and> pathstart h = pathstart \<gamma> \<and> pathfinish h = pathfinish \<gamma> \<and>
       
  3477                           (\<forall>t \<in> {0..1}. norm(h t - \<gamma> t) < d/2)"
       
  3478     using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "d/2"] d by auto
       
  3479   define nn where "nn = 1/(2* pi*\<i>) * contour_integral h (\<lambda>w. 1/(w - z))"
       
  3480   have "\<exists>n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
       
  3481                         pathstart p = pathstart \<gamma> \<and>  pathfinish p = pathfinish \<gamma> \<and>
       
  3482                         (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
       
  3483                         contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
       
  3484                     (is "\<exists>n. \<forall>e > 0. ?PP e n")
       
  3485     proof (rule_tac x=nn in exI, clarify)
       
  3486       fix e::real
       
  3487       assume e: "e>0"
       
  3488       obtain p where p: "polynomial_function p \<and>
       
  3489             pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> (\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < min e (d / 2))"
       
  3490         using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "min e (d/2)"] d \<open>0<e\<close> by auto
       
  3491       have "(\<lambda>w. 1 / (w - z)) holomorphic_on UNIV - {z}"
       
  3492         by (auto simp: intro!: holomorphic_intros)
       
  3493       then show "?PP e nn"
       
  3494         apply (rule_tac x=p in exI)
       
  3495         using pi_eq [of h p] h p d
       
  3496         apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def)
       
  3497         done
       
  3498     qed
       
  3499   then show ?thesis
       
  3500     unfolding winding_number_def
       
  3501     apply (rule someI2_ex)
       
  3502     apply (blast intro: \<open>0<e\<close>)
       
  3503     done
       
  3504 qed
       
  3505 
       
  3506 lemma winding_number_unique:
       
  3507   assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3508       and pi:
       
  3509         "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
       
  3510                           pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
       
  3511                           (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
       
  3512                           contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
       
  3513    shows "winding_number \<gamma> z = n"
       
  3514 proof -
       
  3515   have "path_image \<gamma> \<subseteq> UNIV - {z}"
       
  3516     using assms by blast
       
  3517   then obtain e
       
  3518     where e: "e>0"
       
  3519       and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2;
       
  3520                     (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e);
       
  3521                     pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow>
       
  3522                     contour_integral h2 f = contour_integral h1 f"
       
  3523     using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms  by (auto simp: open_delete)
       
  3524   obtain p where p:
       
  3525      "valid_path p \<and> z \<notin> path_image p \<and>
       
  3526       pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
       
  3527       (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
       
  3528       contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
       
  3529     using pi [OF e] by blast
       
  3530   obtain q where q:
       
  3531      "valid_path q \<and> z \<notin> path_image q \<and>
       
  3532       pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and>
       
  3533       (\<forall>t\<in>{0..1}. cmod (\<gamma> t - q t) < e) \<and> contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
       
  3534     using winding_number [OF \<gamma> e] by blast
       
  3535   have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))"
       
  3536     using p by auto
       
  3537   also have "... = contour_integral q (\<lambda>w. 1 / (w - z))"
       
  3538     apply (rule pi_eq)
       
  3539     using p q
       
  3540     by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros)
       
  3541   also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z"
       
  3542     using q by auto
       
  3543   finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" .
       
  3544   then show ?thesis
       
  3545     by simp
       
  3546 qed
       
  3547 
       
  3548 lemma winding_number_unique_loop:
       
  3549   assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3550       and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  3551       and pi:
       
  3552         "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
       
  3553                            pathfinish p = pathstart p \<and>
       
  3554                            (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
       
  3555                            contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
       
  3556    shows "winding_number \<gamma> z = n"
       
  3557 proof -
       
  3558   have "path_image \<gamma> \<subseteq> UNIV - {z}"
       
  3559     using assms by blast
       
  3560   then obtain e
       
  3561     where e: "e>0"
       
  3562       and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2;
       
  3563                     (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e);
       
  3564                     pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow>
       
  3565                     contour_integral h2 f = contour_integral h1 f"
       
  3566     using contour_integral_nearby_loops [of "UNIV - {z}" \<gamma>] assms  by (auto simp: open_delete)
       
  3567   obtain p where p:
       
  3568      "valid_path p \<and> z \<notin> path_image p \<and>
       
  3569       pathfinish p = pathstart p \<and>
       
  3570       (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
       
  3571       contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
       
  3572     using pi [OF e] by blast
       
  3573   obtain q where q:
       
  3574      "valid_path q \<and> z \<notin> path_image q \<and>
       
  3575       pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and>
       
  3576       (\<forall>t\<in>{0..1}. cmod (\<gamma> t - q t) < e) \<and> contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
       
  3577     using winding_number [OF \<gamma> e] by blast
       
  3578   have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))"
       
  3579     using p by auto
       
  3580   also have "... = contour_integral q (\<lambda>w. 1 / (w - z))"
       
  3581     apply (rule pi_eq)
       
  3582     using p q loop
       
  3583     by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros)
       
  3584   also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z"
       
  3585     using q by auto
       
  3586   finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" .
       
  3587   then show ?thesis
       
  3588     by simp
       
  3589 qed
       
  3590 
       
  3591 lemma winding_number_valid_path:
       
  3592   assumes "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3593     shows "winding_number \<gamma> z = 1/(2*pi*\<i>) * contour_integral \<gamma> (\<lambda>w. 1/(w - z))"
       
  3594   using assms by (auto simp: valid_path_imp_path intro!: winding_number_unique)
       
  3595 
       
  3596 lemma has_contour_integral_winding_number:
       
  3597   assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3598     shows "((\<lambda>w. 1/(w - z)) has_contour_integral (2*pi*\<i>*winding_number \<gamma> z)) \<gamma>"
       
  3599 by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms)
       
  3600 
       
  3601 lemma winding_number_trivial [simp]: "z \<noteq> a \<Longrightarrow> winding_number(linepath a a) z = 0"
       
  3602   by (simp add: winding_number_valid_path)
       
  3603 
       
  3604 lemma winding_number_subpath_trivial [simp]: "z \<noteq> g x \<Longrightarrow> winding_number (subpath x x g) z = 0"
       
  3605   by (simp add: path_image_subpath winding_number_valid_path)
       
  3606 
       
  3607 lemma winding_number_join:
       
  3608   assumes g1: "path g1" "z \<notin> path_image g1"
       
  3609       and g2: "path g2" "z \<notin> path_image g2"
       
  3610       and "pathfinish g1 = pathstart g2"
       
  3611     shows "winding_number(g1 +++ g2) z = winding_number g1 z + winding_number g2 z"
       
  3612   apply (rule winding_number_unique)
       
  3613   using assms apply (simp_all add: not_in_path_image_join)
       
  3614   apply (frule winding_number [OF g2])
       
  3615   apply (frule winding_number [OF g1], clarify)
       
  3616   apply (rename_tac p2 p1)
       
  3617   apply (rule_tac x="p1+++p2" in exI)
       
  3618   apply (simp add: not_in_path_image_join contour_integrable_inversediff algebra_simps)
       
  3619   apply (auto simp: joinpaths_def)
       
  3620   done
       
  3621 
       
  3622 lemma winding_number_reversepath:
       
  3623   assumes "path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3624     shows "winding_number(reversepath \<gamma>) z = - (winding_number \<gamma> z)"
       
  3625   apply (rule winding_number_unique)
       
  3626   using assms
       
  3627   apply simp_all
       
  3628   apply (frule winding_number [OF assms], clarify)
       
  3629   apply (rule_tac x="reversepath p" in exI)
       
  3630   apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse)
       
  3631   apply (auto simp: reversepath_def)
       
  3632   done
       
  3633 
       
  3634 lemma winding_number_shiftpath:
       
  3635   assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3636       and "pathfinish \<gamma> = pathstart \<gamma>" "a \<in> {0..1}"
       
  3637     shows "winding_number(shiftpath a \<gamma>) z = winding_number \<gamma> z"
       
  3638   apply (rule winding_number_unique_loop)
       
  3639   using assms
       
  3640   apply (simp_all add: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath)
       
  3641   apply (frule winding_number [OF \<gamma>], clarify)
       
  3642   apply (rule_tac x="shiftpath a p" in exI)
       
  3643   apply (simp add: contour_integral_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath valid_path_shiftpath)
       
  3644   apply (auto simp: shiftpath_def)
       
  3645   done
       
  3646 
       
  3647 lemma winding_number_split_linepath:
       
  3648   assumes "c \<in> closed_segment a b" "z \<notin> closed_segment a b"
       
  3649     shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z"
       
  3650 proof -
       
  3651   have "z \<notin> closed_segment a c" "z \<notin> closed_segment c b"
       
  3652     using assms  apply (meson convex_contains_segment convex_segment ends_in_segment(1) subsetCE)
       
  3653     using assms  by (meson convex_contains_segment convex_segment ends_in_segment(2) subsetCE)
       
  3654   then show ?thesis
       
  3655     using assms
       
  3656     by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps)
       
  3657 qed
       
  3658 
       
  3659 lemma winding_number_cong:
       
  3660    "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> p t = q t) \<Longrightarrow> winding_number p z = winding_number q z"
       
  3661   by (simp add: winding_number_def pathstart_def pathfinish_def)
       
  3662 
       
  3663 lemma winding_number_offset: "winding_number p z = winding_number (\<lambda>w. p w - z) 0"
       
  3664   apply (simp add: winding_number_def contour_integral_integral path_image_def valid_path_def pathstart_def pathfinish_def)
       
  3665   apply (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe)
       
  3666   apply (rename_tac g)
       
  3667   apply (rule_tac x="\<lambda>t. g t - z" in exI)
       
  3668   apply (force simp: vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise)
       
  3669   apply (rename_tac g)
       
  3670   apply (rule_tac x="\<lambda>t. g t + z" in exI)
       
  3671   apply (simp add: piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise)
       
  3672   apply (force simp: algebra_simps)
       
  3673   done
       
  3674 
       
  3675 (* A combined theorem deducing several things piecewise.*)
       
  3676 lemma winding_number_join_pos_combined:
       
  3677      "\<lbrakk>valid_path \<gamma>1; z \<notin> path_image \<gamma>1; 0 < Re(winding_number \<gamma>1 z);
       
  3678        valid_path \<gamma>2; z \<notin> path_image \<gamma>2; 0 < Re(winding_number \<gamma>2 z); pathfinish \<gamma>1 = pathstart \<gamma>2\<rbrakk>
       
  3679       \<Longrightarrow> valid_path(\<gamma>1 +++ \<gamma>2) \<and> z \<notin> path_image(\<gamma>1 +++ \<gamma>2) \<and> 0 < Re(winding_number(\<gamma>1 +++ \<gamma>2) z)"
       
  3680   by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path)
       
  3681 
       
  3682 
       
  3683 (* Useful sufficient conditions for the winding number to be positive etc.*)
       
  3684 
       
  3685 lemma Re_winding_number:
       
  3686     "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>\<rbrakk>
       
  3687      \<Longrightarrow> Re(winding_number \<gamma> z) = Im(contour_integral \<gamma> (\<lambda>w. 1/(w - z))) / (2*pi)"
       
  3688 by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square)
       
  3689 
       
  3690 lemma winding_number_pos_le:
       
  3691   assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3692       and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> 0 \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))"
       
  3693     shows "0 \<le> Re(winding_number \<gamma> z)"
       
  3694 proof -
       
  3695   have *: "0 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))" if x: "0 < x" "x < 1" for x
       
  3696     using ge by (simp add: Complex.Im_divide algebra_simps x)
       
  3697   show ?thesis
       
  3698     apply (simp add: Re_winding_number [OF \<gamma>] field_simps)
       
  3699     apply (rule has_integral_component_nonneg
       
  3700              [of \<i> "\<lambda>x. if x \<in> {0<..<1}
       
  3701                          then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else 0", simplified])
       
  3702       prefer 3 apply (force simp: *)
       
  3703      apply (simp add: Basis_complex_def)
       
  3704     apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)"])
       
  3705     apply simp
       
  3706     apply (simp only: box_real)
       
  3707     apply (subst has_contour_integral [symmetric])
       
  3708     using \<gamma>
       
  3709     apply (simp add: contour_integrable_inversediff has_contour_integral_integral)
       
  3710     done
       
  3711 qed
       
  3712 
       
  3713 lemma winding_number_pos_lt_lemma:
       
  3714   assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3715       and e: "0 < e"
       
  3716       and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
       
  3717     shows "0 < Re(winding_number \<gamma> z)"
       
  3718 proof -
       
  3719   have "e \<le> Im (contour_integral \<gamma> (\<lambda>w. 1 / (w - z)))"
       
  3720     apply (rule has_integral_component_le
       
  3721              [of \<i> "\<lambda>x. \<i>*e" "\<i>*e" "{0..1}"
       
  3722                     "\<lambda>x. if x \<in> {0<..<1} then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else \<i>*e"
       
  3723                     "contour_integral \<gamma> (\<lambda>w. 1/(w - z))", simplified])
       
  3724     using e
       
  3725     apply (simp_all add: Basis_complex_def)
       
  3726     using has_integral_const_real [of _ 0 1] apply force
       
  3727     apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)", simplified box_real])
       
  3728     apply simp
       
  3729     apply (subst has_contour_integral [symmetric])
       
  3730     using \<gamma>
       
  3731     apply (simp_all add: contour_integrable_inversediff has_contour_integral_integral ge)
       
  3732     done
       
  3733   with e show ?thesis
       
  3734     by (simp add: Re_winding_number [OF \<gamma>] field_simps)
       
  3735 qed
       
  3736 
       
  3737 lemma winding_number_pos_lt:
       
  3738   assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
       
  3739       and e: "0 < e"
       
  3740       and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))"
       
  3741     shows "0 < Re (winding_number \<gamma> z)"
       
  3742 proof -
       
  3743   have bm: "bounded ((\<lambda>w. w - z) ` (path_image \<gamma>))"
       
  3744     using bounded_translation [of _ "-z"] \<gamma> by (simp add: bounded_valid_path_image)
       
  3745   then obtain B where B: "B > 0" and Bno: "\<And>x. x \<in> (\<lambda>w. w - z) ` (path_image \<gamma>) \<Longrightarrow> norm x \<le> B"
       
  3746     using bounded_pos [THEN iffD1, OF bm] by blast
       
  3747   { fix x::real  assume x: "0 < x" "x < 1"
       
  3748     then have B2: "cmod (\<gamma> x - z)^2 \<le> B^2" using Bno [of "\<gamma> x - z"]
       
  3749       by (simp add: path_image_def power2_eq_square mult_mono')
       
  3750     with x have "\<gamma> x \<noteq> z" using \<gamma>
       
  3751       using path_image_def by fastforce
       
  3752     then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) * cnj (\<gamma> x - z)) / (cmod (\<gamma> x - z))\<^sup>2"
       
  3753       using B ge [OF x] B2 e
       
  3754       apply (rule_tac y="e / (cmod (\<gamma> x - z))\<^sup>2" in order_trans)
       
  3755       apply (auto simp: divide_left_mono divide_right_mono)
       
  3756       done
       
  3757     then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
       
  3758       by (simp add: Im_divide_Reals complex_div_cnj [of _ "\<gamma> x - z" for x] del: complex_cnj_diff times_complex.sel)
       
  3759   } note * = this
       
  3760   show ?thesis
       
  3761     using e B by (simp add: * winding_number_pos_lt_lemma [OF \<gamma>, of "e/B^2"])
       
  3762 qed
       
  3763 
       
  3764 subsection\<open>The winding number is an integer\<close>
       
  3765 
       
  3766 text\<open>Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1,
       
  3767      Also on page 134 of Serge Lang's book with the name title, etc.\<close>
       
  3768 
       
  3769 lemma exp_fg:
       
  3770   fixes z::complex
       
  3771   assumes g: "(g has_vector_derivative g') (at x within s)"
       
  3772       and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)"
       
  3773       and z: "g x \<noteq> z"
       
  3774     shows "((\<lambda>x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)"
       
  3775 proof -
       
  3776   have *: "(exp o (\<lambda>x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)"
       
  3777     using assms unfolding has_vector_derivative_def scaleR_conv_of_real
       
  3778     by (auto intro!: derivative_eq_intros)
       
  3779   show ?thesis
       
  3780     apply (rule has_vector_derivative_eq_rhs)
       
  3781     apply (rule bounded_bilinear.has_vector_derivative [OF bounded_bilinear_mult])
       
  3782     using z
       
  3783     apply (auto simp: intro!: derivative_eq_intros * [unfolded o_def] g)
       
  3784     done
       
  3785 qed
       
  3786 
       
  3787 lemma winding_number_exp_integral:
       
  3788   fixes z::complex
       
  3789   assumes \<gamma>: "\<gamma> piecewise_C1_differentiable_on {a..b}"
       
  3790       and ab: "a \<le> b"
       
  3791       and z: "z \<notin> \<gamma> ` {a..b}"
       
  3792     shows "(\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)) integrable_on {a..b}"
       
  3793           (is "?thesis1")
       
  3794           "exp (- (integral {a..b} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))) * (\<gamma> b - z) = \<gamma> a - z"
       
  3795           (is "?thesis2")
       
  3796 proof -
       
  3797   let ?D\<gamma> = "\<lambda>x. vector_derivative \<gamma> (at x)"
       
  3798   have [simp]: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<gamma> x \<noteq> z"
       
  3799     using z by force
       
  3800   have cong: "continuous_on {a..b} \<gamma>"
       
  3801     using \<gamma> by (simp add: piecewise_C1_differentiable_on_def)
       
  3802   obtain k where fink: "finite k" and g_C1_diff: "\<gamma> C1_differentiable_on ({a..b} - k)"
       
  3803     using \<gamma> by (force simp: piecewise_C1_differentiable_on_def)
       
  3804   have o: "open ({a<..<b} - k)"
       
  3805     using \<open>finite k\<close> by (simp add: finite_imp_closed open_Diff)
       
  3806   moreover have "{a<..<b} - k \<subseteq> {a..b} - k"
       
  3807     by force
       
  3808   ultimately have g_diff_at: "\<And>x. \<lbrakk>x \<notin> k; x \<in> {a<..<b}\<rbrakk> \<Longrightarrow> \<gamma> differentiable at x"
       
  3809     by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def differentiable_within_open)
       
  3810   { fix w
       
  3811     assume "w \<noteq> z"
       
  3812     have "continuous_on (ball w (cmod (w - z))) (\<lambda>w. 1 / (w - z))"
       
  3813       by (auto simp: dist_norm intro!: continuous_intros)
       
  3814     moreover have "\<And>x. cmod (w - x) < cmod (w - z) \<Longrightarrow> \<exists>f'. ((\<lambda>w. 1 / (w - z)) has_field_derivative f') (at x)"
       
  3815       by (auto simp: intro!: derivative_eq_intros)
       
  3816     ultimately have "\<exists>h. \<forall>y. norm(y - w) < norm(w - z) \<longrightarrow> (h has_field_derivative 1/(y - z)) (at y)"
       
  3817       using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\<lambda>w. 1/(w - z)"]
       
  3818       by (simp add: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute)
       
  3819   }
       
  3820   then obtain h where h: "\<And>w y. w \<noteq> z \<Longrightarrow> norm(y - w) < norm(w - z) \<Longrightarrow> (h w has_field_derivative 1/(y - z)) (at y)"
       
  3821     by meson
       
  3822   have exy: "\<exists>y. ((\<lambda>x. inverse (\<gamma> x - z) * ?D\<gamma> x) has_integral y) {a..b}"
       
  3823     unfolding integrable_on_def [symmetric]
       
  3824     apply (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \<gamma>], of "-{z}"])
       
  3825     apply (rename_tac w)
       
  3826     apply (rule_tac x="norm(w - z)" in exI)
       
  3827     apply (simp_all add: inverse_eq_divide)
       
  3828     apply (metis has_field_derivative_at_within h)
       
  3829     done
       
  3830   have vg_int: "(\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)) integrable_on {a..b}"
       
  3831     unfolding box_real [symmetric] divide_inverse_commute
       
  3832     by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab)
       
  3833   with ab show ?thesis1
       
  3834     by (simp add: divide_inverse_commute integral_def integrable_on_def)
       
  3835   { fix t
       
  3836     assume t: "t \<in> {a..b}"
       
  3837     have cball: "continuous_on (ball (\<gamma> t) (dist (\<gamma> t) z)) (\<lambda>x. inverse (x - z))"
       
  3838         using z by (auto intro!: continuous_intros simp: dist_norm)
       
  3839     have icd: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow> (\<lambda>w. inverse (w - z)) field_differentiable at x"
       
  3840       unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros)
       
  3841     obtain h where h: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow>
       
  3842                        (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\<gamma> t - y) < cmod (\<gamma> t - z)})"
       
  3843       using holomorphic_convex_primitive [where f = "\<lambda>w. inverse(w - z)", OF convex_ball finite.emptyI cball icd]
       
  3844       by simp (auto simp: ball_def dist_norm that)
       
  3845     { fix x D
       
  3846       assume x: "x \<notin> k" "a < x" "x < b"
       
  3847       then have "x \<in> interior ({a..b} - k)"
       
  3848         using open_subset_interior [OF o] by fastforce
       
  3849       then have con: "isCont (\<lambda>x. ?D\<gamma> x) x"
       
  3850         using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior)
       
  3851       then have con_vd: "continuous (at x within {a..b}) (\<lambda>x. ?D\<gamma> x)"
       
  3852         by (rule continuous_at_imp_continuous_within)
       
  3853       have gdx: "\<gamma> differentiable at x"
       
  3854         using x by (simp add: g_diff_at)
       
  3855       have "((\<lambda>c. exp (- integral {a..c} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))) * (\<gamma> c - z)) has_derivative (\<lambda>h. 0))
       
  3856           (at x within {a..b})"
       
  3857         using x gdx t
       
  3858         apply (clarsimp simp add: differentiable_iff_scaleR)
       
  3859         apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_within)
       
  3860         apply (simp_all add: has_vector_derivative_def [symmetric])
       
  3861         apply (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at])
       
  3862         apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+
       
  3863         done
       
  3864       } note * = this
       
  3865     have "exp (- (integral {a..t} (\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)))) * (\<gamma> t - z) =\<gamma> a - z"
       
  3866       apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \<union> k" a b])
       
  3867       using t
       
  3868       apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous [OF vg_int]  simp add: ab)+
       
  3869       done
       
  3870    }
       
  3871   with ab show ?thesis2
       
  3872     by (simp add: divide_inverse_commute integral_def)
       
  3873 qed
       
  3874 
       
  3875 corollary winding_number_exp_2pi:
       
  3876     "\<lbrakk>path p; z \<notin> path_image p\<rbrakk>
       
  3877      \<Longrightarrow> pathfinish p - z = exp (2 * pi * \<i> * winding_number p z) * (pathstart p - z)"
       
  3878 using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def
       
  3879   by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus)
       
  3880 
       
  3881 
       
  3882 subsection\<open>The version with complex integers and equality\<close>
       
  3883 
       
  3884 lemma integer_winding_number_eq:
       
  3885   assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>"
       
  3886   shows "winding_number \<gamma> z \<in> \<int> \<longleftrightarrow> pathfinish \<gamma> = pathstart \<gamma>"
       
  3887 proof -
       
  3888   have *: "\<And>i::complex. \<And>g0 g1. \<lbrakk>i \<noteq> 0; g0 \<noteq> z; (g1 - z) / i = g0 - z\<rbrakk> \<Longrightarrow> (i = 1 \<longleftrightarrow> g1 = g0)"
       
  3889       by (simp add: field_simps) algebra
       
  3890   obtain p where p: "valid_path p" "z \<notin> path_image p"
       
  3891                     "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>"
       
  3892                     "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
       
  3893     using winding_number [OF assms, of 1] by auto
       
  3894   have [simp]: "(winding_number \<gamma> z \<in> \<int>) = (exp (contour_integral p (\<lambda>w. 1 / (w - z))) = 1)"
       
  3895       using p by (simp add: exp_eq_1 complex_is_Int_iff)
       
  3896   have "winding_number p z \<in> \<int> \<longleftrightarrow> pathfinish p = pathstart p"
       
  3897     using p z
       
  3898     apply (simp add: winding_number_valid_path valid_path_def path_image_def pathstart_def pathfinish_def)
       
  3899     using winding_number_exp_integral(2) [of p 0 1 z]
       
  3900     apply (simp add: field_simps contour_integral_integral exp_minus)
       
  3901     apply (rule *)
       
  3902     apply (auto simp: path_image_def field_simps)
       
  3903     done
       
  3904   then show ?thesis using p
       
  3905     by (auto simp: winding_number_valid_path)
       
  3906 qed
       
  3907 
       
  3908 theorem integer_winding_number:
       
  3909   "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> path_image \<gamma>\<rbrakk> \<Longrightarrow> winding_number \<gamma> z \<in> \<int>"
       
  3910 by (metis integer_winding_number_eq)
       
  3911 
       
  3912 
       
  3913 text\<open>If the winding number's magnitude is at least one, then the path must contain points in every direction.*)
       
  3914    We can thus bound the winding number of a path that doesn't intersect a given ray. \<close>
       
  3915 
       
  3916 lemma winding_number_pos_meets:
       
  3917   fixes z::complex
       
  3918   assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and 1: "Re (winding_number \<gamma> z) \<ge> 1"
       
  3919       and w: "w \<noteq> z"
       
  3920   shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>"
       
  3921 proof -
       
  3922   have [simp]: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<gamma> x \<noteq> z"
       
  3923     using z by (auto simp: path_image_def)
       
  3924   have [simp]: "z \<notin> \<gamma> ` {0..1}"
       
  3925     using path_image_def z by auto
       
  3926   have gpd: "\<gamma> piecewise_C1_differentiable_on {0..1}"
       
  3927     using \<gamma> valid_path_def by blast
       
  3928   define r where "r = (w - z) / (\<gamma> 0 - z)"
       
  3929   have [simp]: "r \<noteq> 0"
       
  3930     using w z by (auto simp: r_def)
       
  3931   have "Arg r \<le> 2*pi"
       
  3932     by (simp add: Arg less_eq_real_def)
       
  3933   also have "... \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))"
       
  3934     using 1
       
  3935     apply (simp add: winding_number_valid_path [OF \<gamma> z] Cauchy_Integral_Thm.contour_integral_integral)
       
  3936     apply (simp add: Complex.Re_divide field_simps power2_eq_square)
       
  3937     done
       
  3938   finally have "Arg r \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))" .
       
  3939   then have "\<exists>t. t \<in> {0..1} \<and> Im(integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r"
       
  3940     apply (simp add:)
       
  3941     apply (rule Topological_Spaces.IVT')
       
  3942     apply (simp_all add: Complex_Transcendental.Arg_ge_0)
       
  3943     apply (intro continuous_intros indefinite_integral_continuous winding_number_exp_integral [OF gpd]; simp)
       
  3944     done
       
  3945   then obtain t where t:     "t \<in> {0..1}"
       
  3946                   and eqArg: "Im (integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r"
       
  3947     by blast
       
  3948   define i where "i = integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
       
  3949   have iArg: "Arg r = Im i"
       
  3950     using eqArg by (simp add: i_def)
       
  3951   have gpdt: "\<gamma> piecewise_C1_differentiable_on {0..t}"
       
  3952     by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t)
       
  3953   have "exp (- i) * (\<gamma> t - z) = \<gamma> 0 - z"
       
  3954     unfolding i_def
       
  3955     apply (rule winding_number_exp_integral [OF gpdt])
       
  3956     using t z unfolding path_image_def
       
  3957     apply force+
       
  3958     done
       
  3959   then have *: "\<gamma> t - z = exp i * (\<gamma> 0 - z)"
       
  3960     by (simp add: exp_minus field_simps)
       
  3961   then have "(w - z) = r * (\<gamma> 0 - z)"
       
  3962     by (simp add: r_def)
       
  3963   then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \<gamma> t"
       
  3964     apply (simp add:)
       
  3965     apply (subst Complex_Transcendental.Arg_eq [of r])
       
  3966     apply (simp add: iArg)
       
  3967     using *
       
  3968     apply (simp add: exp_eq_polar field_simps)
       
  3969     done
       
  3970   with t show ?thesis
       
  3971     by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def)
       
  3972 qed
       
  3973 
       
  3974 lemma winding_number_big_meets:
       
  3975   fixes z::complex
       
  3976   assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "\<bar>Re (winding_number \<gamma> z)\<bar> \<ge> 1"
       
  3977       and w: "w \<noteq> z"
       
  3978   shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>"
       
  3979 proof -
       
  3980   { assume "Re (winding_number \<gamma> z) \<le> - 1"
       
  3981     then have "Re (winding_number (reversepath \<gamma>) z) \<ge> 1"
       
  3982       by (simp add: \<gamma> valid_path_imp_path winding_number_reversepath z)
       
  3983     moreover have "valid_path (reversepath \<gamma>)"
       
  3984       using \<gamma> valid_path_imp_reverse by auto
       
  3985     moreover have "z \<notin> path_image (reversepath \<gamma>)"
       
  3986       by (simp add: z)
       
  3987     ultimately have "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image (reversepath \<gamma>)"
       
  3988       using winding_number_pos_meets w by blast
       
  3989     then have ?thesis
       
  3990       by simp
       
  3991   }
       
  3992   then show ?thesis
       
  3993     using assms
       
  3994     by (simp add: Groups.abs_if_class.abs_if winding_number_pos_meets split: if_split_asm)
       
  3995 qed
       
  3996 
       
  3997 lemma winding_number_less_1:
       
  3998   fixes z::complex
       
  3999   shows
       
  4000   "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>; w \<noteq> z;
       
  4001     \<And>a::real. 0 < a \<Longrightarrow> z + a*(w - z) \<notin> path_image \<gamma>\<rbrakk>
       
  4002    \<Longrightarrow> \<bar>Re(winding_number \<gamma> z)\<bar> < 1"
       
  4003    by (auto simp: not_less dest: winding_number_big_meets)
       
  4004 
       
  4005 text\<open>One way of proving that WN=1 for a loop.\<close>
       
  4006 lemma winding_number_eq_1:
       
  4007   fixes z::complex
       
  4008   assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  4009       and 0: "0 < Re(winding_number \<gamma> z)" and 2: "Re(winding_number \<gamma> z) < 2"
       
  4010   shows "winding_number \<gamma> z = 1"
       
  4011 proof -
       
  4012   have "winding_number \<gamma> z \<in> Ints"
       
  4013     by (simp add: \<gamma> integer_winding_number loop valid_path_imp_path z)
       
  4014   then show ?thesis
       
  4015     using 0 2 by (auto simp: Ints_def)
       
  4016 qed
       
  4017 
       
  4018 
       
  4019 subsection\<open>Continuity of winding number and invariance on connected sets.\<close>
       
  4020 
       
  4021 lemma continuous_at_winding_number:
       
  4022   fixes z::complex
       
  4023   assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>"
       
  4024   shows "continuous (at z) (winding_number \<gamma>)"
       
  4025 proof -
       
  4026   obtain e where "e>0" and cbg: "cball z e \<subseteq> - path_image \<gamma>"
       
  4027     using open_contains_cball [of "- path_image \<gamma>"]  z
       
  4028     by (force simp: closed_def [symmetric] closed_path_image [OF \<gamma>])
       
  4029   then have ppag: "path_image \<gamma> \<subseteq> - cball z (e/2)"
       
  4030     by (force simp: cball_def dist_norm)
       
  4031   have oc: "open (- cball z (e / 2))"
       
  4032     by (simp add: closed_def [symmetric])
       
  4033   obtain d where "d>0" and pi_eq:
       
  4034     "\<And>h1 h2. \<lbrakk>valid_path h1; valid_path h2;
       
  4035               (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d);
       
  4036               pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\<rbrakk>
       
  4037              \<Longrightarrow>
       
  4038                path_image h1 \<subseteq> - cball z (e / 2) \<and>
       
  4039                path_image h2 \<subseteq> - cball z (e / 2) \<and>
       
  4040                (\<forall>f. f holomorphic_on - cball z (e / 2) \<longrightarrow> contour_integral h2 f = contour_integral h1 f)"
       
  4041     using contour_integral_nearby_ends [OF oc \<gamma> ppag] by metis
       
  4042   obtain p where p: "valid_path p" "z \<notin> path_image p"
       
  4043                     "pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma>"
       
  4044               and pg: "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (\<gamma> t - p t) < min d e / 2"
       
  4045               and pi: "contour_integral p (\<lambda>x. 1 / (x - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
       
  4046     using winding_number [OF \<gamma> z, of "min d e / 2"] \<open>d>0\<close> \<open>e>0\<close> by auto
       
  4047   { fix w
       
  4048     assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2"
       
  4049     then have wnotp: "w \<notin> path_image p"
       
  4050       using cbg \<open>d>0\<close> \<open>e>0\<close>
       
  4051       apply (simp add: path_image_def cball_def dist_norm, clarify)
       
  4052       apply (frule pg)
       
  4053       apply (drule_tac c="\<gamma> x" in subsetD)
       
  4054       apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l)
       
  4055       done
       
  4056     have wnotg: "w \<notin> path_image \<gamma>"
       
  4057       using cbg e2 \<open>e>0\<close> by (force simp: dist_norm norm_minus_commute)
       
  4058     { fix k::real
       
  4059       assume k: "k>0"
       
  4060       then obtain q where q: "valid_path q" "w \<notin> path_image q"
       
  4061                              "pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma>"
       
  4062                     and qg: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (\<gamma> t - q t) < min k (min d e) / 2"
       
  4063                     and qi: "contour_integral q (\<lambda>u. 1 / (u - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
       
  4064         using winding_number [OF \<gamma> wnotg, of "min k (min d e) / 2"] \<open>d>0\<close> \<open>e>0\<close> k
       
  4065         by (force simp: min_divide_distrib_right)
       
  4066       have "contour_integral p (\<lambda>u. 1 / (u - w)) = contour_integral q (\<lambda>u. 1 / (u - w))"
       
  4067         apply (rule pi_eq [OF \<open>valid_path q\<close> \<open>valid_path p\<close>, THEN conjunct2, THEN conjunct2, rule_format])
       
  4068         apply (frule pg)
       
  4069         apply (frule qg)
       
  4070         using p q \<open>d>0\<close> e2
       
  4071         apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
       
  4072         done
       
  4073       then have "contour_integral p (\<lambda>x. 1 / (x - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
       
  4074         by (simp add: pi qi)
       
  4075     } note pip = this
       
  4076     have "path p"
       
  4077       using p by (simp add: valid_path_imp_path)
       
  4078     then have "winding_number p w = winding_number \<gamma> w"
       
  4079       apply (rule winding_number_unique [OF _ wnotp])
       
  4080       apply (rule_tac x=p in exI)
       
  4081       apply (simp add: p wnotp min_divide_distrib_right pip)
       
  4082       done
       
  4083   } note wnwn = this
       
  4084   obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \<subseteq> - path_image p"
       
  4085     using p open_contains_cball [of "- path_image p"]
       
  4086     by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path])
       
  4087   obtain L
       
  4088     where "L>0"
       
  4089       and L: "\<And>f B. \<lbrakk>f holomorphic_on - cball z (3 / 4 * pe);
       
  4090                       \<forall>z \<in> - cball z (3 / 4 * pe). cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
       
  4091                       cmod (contour_integral p f) \<le> L * B"
       
  4092     using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \<open>valid_path p\<close> by force
       
  4093   { fix e::real and w::complex
       
  4094     assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)"
       
  4095     then have [simp]: "w \<notin> path_image p"
       
  4096       using cbp p(2) \<open>0 < pe\<close>
       
  4097       by (force simp: dist_norm norm_minus_commute path_image_def cball_def)
       
  4098     have [simp]: "contour_integral p (\<lambda>x. 1/(x - w)) - contour_integral p (\<lambda>x. 1/(x - z)) =
       
  4099                   contour_integral p (\<lambda>x. 1/(x - w) - 1/(x - z))"
       
  4100       by (simp add: p contour_integrable_inversediff contour_integral_diff)
       
  4101     { fix x
       
  4102       assume pe: "3/4 * pe < cmod (z - x)"
       
  4103       have "cmod (w - x) < pe/4 + cmod (z - x)"
       
  4104         by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1))
       
  4105       then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp
       
  4106       have "cmod (z - x) \<le> cmod (z - w) + cmod (w - x)"
       
  4107         using norm_diff_triangle_le by blast
       
  4108       also have "... < pe/4 + cmod (w - x)"
       
  4109         using w by (simp add: norm_minus_commute)
       
  4110       finally have "pe/2 < cmod (w - x)"
       
  4111         using pe by auto
       
  4112       then have "(pe/2)^2 < cmod (w - x) ^ 2"
       
  4113         apply (rule power_strict_mono)
       
  4114         using \<open>pe>0\<close> by auto
       
  4115       then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2"
       
  4116         by (simp add: power_divide)
       
  4117       have "8 * L * cmod (w - z) < e * pe\<^sup>2"
       
  4118         using w \<open>L>0\<close> by (simp add: field_simps)
       
  4119       also have "... < e * 4 * cmod (w - x) * cmod (w - x)"
       
  4120         using pe2 \<open>e>0\<close> by (simp add: power2_eq_square)
       
  4121       also have "... < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))"
       
  4122         using wx
       
  4123         apply (rule mult_strict_left_mono)
       
  4124         using pe2 e not_less_iff_gr_or_eq by fastforce
       
  4125       finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)"
       
  4126         by simp
       
  4127       also have "... \<le> e * cmod (w - x) * cmod (z - x)"
       
  4128          using e by simp
       
  4129       finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" .
       
  4130       have "L * cmod (1 / (x - w) - 1 / (x - z)) \<le> e"
       
  4131         apply (cases "x=z \<or> x=w")
       
  4132         using pe \<open>pe>0\<close> w \<open>L>0\<close>
       
  4133         apply (force simp: norm_minus_commute)
       
  4134         using wx w(2) \<open>L>0\<close> pe pe2 Lwz
       
  4135         apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square)
       
  4136         done
       
  4137     } note L_cmod_le = this
       
  4138     have *: "cmod (contour_integral p (\<lambda>x. 1 / (x - w) - 1 / (x - z))) \<le> L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)"
       
  4139       apply (rule L)
       
  4140       using \<open>pe>0\<close> w
       
  4141       apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
       
  4142       using \<open>pe>0\<close> w \<open>L>0\<close>
       
  4143       apply (auto simp: cball_def dist_norm field_simps L_cmod_le  simp del: less_divide_eq_numeral1 le_divide_eq_numeral1)
       
  4144       done
       
  4145     have "cmod (contour_integral p (\<lambda>x. 1 / (x - w)) - contour_integral p (\<lambda>x. 1 / (x - z))) < 2*e"
       
  4146       apply (simp add:)
       
  4147       apply (rule le_less_trans [OF *])
       
  4148       using \<open>L>0\<close> e
       
  4149       apply (force simp: field_simps)
       
  4150       done
       
  4151     then have "cmod (winding_number p w - winding_number p z) < e"
       
  4152       using pi_ge_two e
       
  4153       by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans)
       
  4154   } note cmod_wn_diff = this
       
  4155   then have "isCont (winding_number p) z"
       
  4156     apply (simp add: continuous_at_eps_delta, clarify)
       
  4157     apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI)
       
  4158     using \<open>pe>0\<close> \<open>L>0\<close>
       
  4159     apply (simp add: dist_norm cmod_wn_diff)
       
  4160     done
       
  4161   then show ?thesis
       
  4162     apply (rule continuous_transform_within [where d = "min d e / 2"])
       
  4163     apply (auto simp: \<open>d>0\<close> \<open>e>0\<close> dist_norm wnwn)
       
  4164     done
       
  4165 qed
       
  4166 
       
  4167 corollary continuous_on_winding_number:
       
  4168     "path \<gamma> \<Longrightarrow> continuous_on (- path_image \<gamma>) (\<lambda>w. winding_number \<gamma> w)"
       
  4169   by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number)
       
  4170 
       
  4171 
       
  4172 subsection\<open>The winding number is constant on a connected region\<close>
       
  4173 
       
  4174 lemma winding_number_constant:
       
  4175   assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and cs: "connected s" and sg: "s \<inter> path_image \<gamma> = {}"
       
  4176     shows "\<exists>k. \<forall>z \<in> s. winding_number \<gamma> z = k"
       
  4177 proof -
       
  4178   { fix y z
       
  4179     assume ne: "winding_number \<gamma> y \<noteq> winding_number \<gamma> z"
       
  4180     assume "y \<in> s" "z \<in> s"
       
  4181     then have "winding_number \<gamma> y \<in> \<int>"  "winding_number \<gamma> z \<in>  \<int>"
       
  4182       using integer_winding_number [OF \<gamma> loop] sg \<open>y \<in> s\<close> by auto
       
  4183     with ne have "1 \<le> cmod (winding_number \<gamma> y - winding_number \<gamma> z)"
       
  4184       by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff)
       
  4185   } note * = this
       
  4186   show ?thesis
       
  4187     apply (rule continuous_discrete_range_constant [OF cs])
       
  4188     using continuous_on_winding_number [OF \<gamma>] sg
       
  4189     apply (metis Diff_Compl Diff_eq_empty_iff continuous_on_subset)
       
  4190     apply (rule_tac x=1 in exI)
       
  4191     apply (auto simp: *)
       
  4192     done
       
  4193 qed
       
  4194 
       
  4195 lemma winding_number_eq:
       
  4196      "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; w \<in> s; z \<in> s; connected s; s \<inter> path_image \<gamma> = {}\<rbrakk>
       
  4197       \<Longrightarrow> winding_number \<gamma> w = winding_number \<gamma> z"
       
  4198 using winding_number_constant by fastforce
       
  4199 
       
  4200 lemma open_winding_number_levelsets:
       
  4201   assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  4202     shows "open {z. z \<notin> path_image \<gamma> \<and> winding_number \<gamma> z = k}"
       
  4203 proof -
       
  4204   have op: "open (- path_image \<gamma>)"
       
  4205     by (simp add: closed_path_image \<gamma> open_Compl)
       
  4206   { fix z assume z: "z \<notin> path_image \<gamma>" and k: "k = winding_number \<gamma> z"
       
  4207     obtain e where e: "e>0" "ball z e \<subseteq> - path_image \<gamma>"
       
  4208       using open_contains_ball [of "- path_image \<gamma>"] op z
       
  4209       by blast
       
  4210     have "\<exists>e>0. \<forall>y. dist y z < e \<longrightarrow> y \<notin> path_image \<gamma> \<and> winding_number \<gamma> y = winding_number \<gamma> z"
       
  4211       apply (rule_tac x=e in exI)
       
  4212       using e apply (simp add: dist_norm ball_def norm_minus_commute)
       
  4213       apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where s = "ball z e"])
       
  4214       done
       
  4215   } then
       
  4216   show ?thesis
       
  4217     by (auto simp: open_dist)
       
  4218 qed
       
  4219 
       
  4220 subsection\<open>Winding number is zero "outside" a curve, in various senses\<close>
       
  4221 
       
  4222 lemma winding_number_zero_in_outside:
       
  4223   assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and z: "z \<in> outside (path_image \<gamma>)"
       
  4224     shows "winding_number \<gamma> z = 0"
       
  4225 proof -
       
  4226   obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B"
       
  4227     using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto
       
  4228   obtain w::complex where w: "w \<notin> ball 0 (B + 1)"
       
  4229     by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real)
       
  4230   have "- ball 0 (B + 1) \<subseteq> outside (path_image \<gamma>)"
       
  4231     apply (rule outside_subset_convex)
       
  4232     using B subset_ball by auto
       
  4233   then have wout: "w \<in> outside (path_image \<gamma>)"
       
  4234     using w by blast
       
  4235   moreover obtain k where "\<And>z. z \<in> outside (path_image \<gamma>) \<Longrightarrow> winding_number \<gamma> z = k"
       
  4236     using winding_number_constant [OF \<gamma> loop, of "outside(path_image \<gamma>)"] connected_outside
       
  4237     by (metis DIM_complex bounded_path_image dual_order.refl \<gamma> outside_no_overlap)
       
  4238   ultimately have "winding_number \<gamma> z = winding_number \<gamma> w"
       
  4239     using z by blast
       
  4240   also have "... = 0"
       
  4241   proof -
       
  4242     have wnot: "w \<notin> path_image \<gamma>"  using wout by (simp add: outside_def)
       
  4243     { fix e::real assume "0<e"
       
  4244       obtain p where p: "polynomial_function p" "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>"
       
  4245                  and pg1: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < 1)"
       
  4246                  and pge: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < e)"
       
  4247         using path_approx_polynomial_function [OF \<gamma>, of "min 1 e"] \<open>e>0\<close> by force
       
  4248       have pip: "path_image p \<subseteq> ball 0 (B + 1)"
       
  4249         using B
       
  4250         apply (clarsimp simp add: path_image_def dist_norm ball_def)
       
  4251         apply (frule (1) pg1)
       
  4252         apply (fastforce dest: norm_add_less)
       
  4253         done
       
  4254       then have "w \<notin> path_image p"  using w by blast
       
  4255       then have "\<exists>p. valid_path p \<and> w \<notin> path_image p \<and>
       
  4256                      pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
       
  4257                      (\<forall>t\<in>{0..1}. cmod (\<gamma> t - p t) < e) \<and> contour_integral p (\<lambda>wa. 1 / (wa - w)) = 0"
       
  4258         apply (rule_tac x=p in exI)
       
  4259         apply (simp add: p valid_path_polynomial_function)
       
  4260         apply (intro conjI)
       
  4261         using pge apply (simp add: norm_minus_commute)
       
  4262         apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]])
       
  4263         apply (rule holomorphic_intros | simp add: dist_norm)+
       
  4264         using mem_ball_0 w apply blast
       
  4265         using p apply (simp_all add: valid_path_polynomial_function loop pip)
       
  4266         done
       
  4267     }
       
  4268     then show ?thesis
       
  4269       by (auto intro: winding_number_unique [OF \<gamma>] simp add: wnot)
       
  4270   qed
       
  4271   finally show ?thesis .
       
  4272 qed
       
  4273 
       
  4274 lemma winding_number_zero_outside:
       
  4275     "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> s; path_image \<gamma> \<subseteq> s\<rbrakk> \<Longrightarrow> winding_number \<gamma> z = 0"
       
  4276   by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside)
       
  4277 
       
  4278 lemma winding_number_zero_at_infinity:
       
  4279   assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  4280     shows "\<exists>B. \<forall>z. B \<le> norm z \<longrightarrow> winding_number \<gamma> z = 0"
       
  4281 proof -
       
  4282   obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B"
       
  4283     using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto
       
  4284   then show ?thesis
       
  4285     apply (rule_tac x="B+1" in exI, clarify)
       
  4286     apply (rule winding_number_zero_outside [OF \<gamma> convex_cball [of 0 B] loop])
       
  4287     apply (meson less_add_one mem_cball_0 not_le order_trans)
       
  4288     using ball_subset_cball by blast
       
  4289 qed
       
  4290 
       
  4291 lemma winding_number_zero_point:
       
  4292     "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; open s; path_image \<gamma> \<subseteq> s\<rbrakk>
       
  4293      \<Longrightarrow> \<exists>z. z \<in> s \<and> winding_number \<gamma> z = 0"
       
  4294   using outside_compact_in_open [of "path_image \<gamma>" s] path_image_nonempty winding_number_zero_in_outside
       
  4295   by (fastforce simp add: compact_path_image)
       
  4296 
       
  4297 
       
  4298 text\<open>If a path winds round a set, it winds rounds its inside.\<close>
       
  4299 lemma winding_number_around_inside:
       
  4300   assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  4301       and cls: "closed s" and cos: "connected s" and s_disj: "s \<inter> path_image \<gamma> = {}"
       
  4302       and z: "z \<in> s" and wn_nz: "winding_number \<gamma> z \<noteq> 0" and w: "w \<in> s \<union> inside s"
       
  4303     shows "winding_number \<gamma> w = winding_number \<gamma> z"
       
  4304 proof -
       
  4305   have ssb: "s \<subseteq> inside(path_image \<gamma>)"
       
  4306   proof
       
  4307     fix x :: complex
       
  4308     assume "x \<in> s"
       
  4309     hence "x \<notin> path_image \<gamma>"
       
  4310       by (meson disjoint_iff_not_equal s_disj)
       
  4311     thus "x \<in> inside (path_image \<gamma>)"
       
  4312       using \<open>x \<in> s\<close> by (metis (no_types) ComplI UnE cos \<gamma> loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z)
       
  4313 qed
       
  4314   show ?thesis
       
  4315     apply (rule winding_number_eq [OF \<gamma> loop w])
       
  4316     using z apply blast
       
  4317     apply (simp add: cls connected_with_inside cos)
       
  4318     apply (simp add: Int_Un_distrib2 s_disj, safe)
       
  4319     by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \<gamma>"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \<gamma> inside_no_overlap)
       
  4320  qed
       
  4321 
       
  4322 
       
  4323 text\<open>Bounding a WN by 1/2 for a path and point in opposite halfspaces.\<close>
       
  4324 lemma winding_number_subpath_continuous:
       
  4325   assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>"
       
  4326     shows "continuous_on {0..1} (\<lambda>x. winding_number(subpath 0 x \<gamma>) z)"
       
  4327 proof -
       
  4328   have *: "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) =
       
  4329          winding_number (subpath 0 x \<gamma>) z"
       
  4330          if x: "0 \<le> x" "x \<le> 1" for x
       
  4331   proof -
       
  4332     have "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) =
       
  4333           1 / (2*pi*\<i>) * contour_integral (subpath 0 x \<gamma>) (\<lambda>w. 1/(w - z))"
       
  4334       using assms x
       
  4335       apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff])
       
  4336       done
       
  4337     also have "... = winding_number (subpath 0 x \<gamma>) z"
       
  4338       apply (subst winding_number_valid_path)
       
  4339       using assms x
       
  4340       apply (simp_all add: path_image_subpath valid_path_subpath)
       
  4341       by (force simp: path_image_def)
       
  4342     finally show ?thesis .
       
  4343   qed
       
  4344   show ?thesis
       
  4345     apply (rule continuous_on_eq
       
  4346                  [where f = "\<lambda>x. 1 / (2*pi*\<i>) *
       
  4347                                  integral {0..x} (\<lambda>t. 1/(\<gamma> t - z) * vector_derivative \<gamma> (at t))"])
       
  4348     apply (rule continuous_intros)+
       
  4349     apply (rule indefinite_integral_continuous)
       
  4350     apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on])
       
  4351       using assms
       
  4352     apply (simp add: *)
       
  4353     done
       
  4354 qed
       
  4355 
       
  4356 lemma winding_number_ivt_pos:
       
  4357     assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "0 \<le> w" "w \<le> Re(winding_number \<gamma> z)"
       
  4358       shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w"
       
  4359   apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right])
       
  4360   apply (simp add:)
       
  4361   apply (rule winding_number_subpath_continuous [OF \<gamma> z])
       
  4362   using assms
       
  4363   apply (auto simp: path_image_def image_def)
       
  4364   done
       
  4365 
       
  4366 lemma winding_number_ivt_neg:
       
  4367     assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "Re(winding_number \<gamma> z) \<le> w" "w \<le> 0"
       
  4368       shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w"
       
  4369   apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right])
       
  4370   apply (simp add:)
       
  4371   apply (rule winding_number_subpath_continuous [OF \<gamma> z])
       
  4372   using assms
       
  4373   apply (auto simp: path_image_def image_def)
       
  4374   done
       
  4375 
       
  4376 lemma winding_number_ivt_abs:
       
  4377     assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "0 \<le> w" "w \<le> \<bar>Re(winding_number \<gamma> z)\<bar>"
       
  4378       shows "\<exists>t \<in> {0..1}. \<bar>Re (winding_number (subpath 0 t \<gamma>) z)\<bar> = w"
       
  4379   using assms winding_number_ivt_pos [of \<gamma> z w] winding_number_ivt_neg [of \<gamma> z "-w"]
       
  4380   by force
       
  4381 
       
  4382 lemma winding_number_lt_half_lemma:
       
  4383   assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and az: "a \<bullet> z \<le> b" and pag: "path_image \<gamma> \<subseteq> {w. a \<bullet> w > b}"
       
  4384     shows "Re(winding_number \<gamma> z) < 1/2"
       
  4385 proof -
       
  4386   { assume "Re(winding_number \<gamma> z) \<ge> 1/2"
       
  4387     then obtain t::real where t: "0 \<le> t" "t \<le> 1" and sub12: "Re (winding_number (subpath 0 t \<gamma>) z) = 1/2"
       
  4388       using winding_number_ivt_pos [OF \<gamma> z, of "1/2"] by auto
       
  4389     have gt: "\<gamma> t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \<gamma>) z)))) * (\<gamma> 0 - z))"
       
  4390       using winding_number_exp_2pi [of "subpath 0 t \<gamma>" z]
       
  4391       apply (simp add: t \<gamma> valid_path_imp_path)
       
  4392       using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12)
       
  4393     have "b < a \<bullet> \<gamma> 0"
       
  4394     proof -
       
  4395       have "\<gamma> 0 \<in> {c. b < a \<bullet> c}"
       
  4396         by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one)
       
  4397       thus ?thesis
       
  4398         by blast
       
  4399     qed
       
  4400     moreover have "b < a \<bullet> \<gamma> t"
       
  4401     proof -
       
  4402       have "\<gamma> t \<in> {c. b < a \<bullet> c}"
       
  4403         by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t)
       
  4404       thus ?thesis
       
  4405         by blast
       
  4406     qed
       
  4407     ultimately have "0 < a \<bullet> (\<gamma> 0 - z)" "0 < a \<bullet> (\<gamma> t - z)" using az
       
  4408       by (simp add: inner_diff_right)+
       
  4409     then have False
       
  4410       by (simp add: gt inner_mult_right mult_less_0_iff)
       
  4411   }
       
  4412   then show ?thesis by force
       
  4413 qed
       
  4414 
       
  4415 lemma winding_number_lt_half:
       
  4416   assumes "valid_path \<gamma>" "a \<bullet> z \<le> b" "path_image \<gamma> \<subseteq> {w. a \<bullet> w > b}"
       
  4417     shows "\<bar>Re (winding_number \<gamma> z)\<bar> < 1/2"
       
  4418 proof -
       
  4419   have "z \<notin> path_image \<gamma>" using assms by auto
       
  4420   with assms show ?thesis
       
  4421     apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1)
       
  4422     apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \<gamma> z a b]
       
  4423                  winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath)
       
  4424     done
       
  4425 qed
       
  4426 
       
  4427 lemma winding_number_le_half:
       
  4428   assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>"
       
  4429       and anz: "a \<noteq> 0" and azb: "a \<bullet> z \<le> b" and pag: "path_image \<gamma> \<subseteq> {w. a \<bullet> w \<ge> b}"
       
  4430     shows "\<bar>Re (winding_number \<gamma> z)\<bar> \<le> 1/2"
       
  4431 proof -
       
  4432   { assume wnz_12: "\<bar>Re (winding_number \<gamma> z)\<bar> > 1/2"
       
  4433     have "isCont (winding_number \<gamma>) z"
       
  4434       by (metis continuous_at_winding_number valid_path_imp_path \<gamma> z)
       
  4435     then obtain d where "d>0" and d: "\<And>x'. dist x' z < d \<Longrightarrow> dist (winding_number \<gamma> x') (winding_number \<gamma> z) < \<bar>Re(winding_number \<gamma> z)\<bar> - 1/2"
       
  4436       using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast
       
  4437     define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a"
       
  4438     have *: "a \<bullet> z' \<le> b - d / 3 * cmod a"
       
  4439       unfolding z'_def inner_mult_right' divide_inverse
       
  4440       apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz)
       
  4441       apply (metis \<open>0 < d\<close> add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral)
       
  4442       done
       
  4443     have "cmod (winding_number \<gamma> z' - winding_number \<gamma> z) < \<bar>Re (winding_number \<gamma> z)\<bar> - 1/2"
       
  4444       using d [of z'] anz \<open>d>0\<close> by (simp add: dist_norm z'_def)
       
  4445     then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - cmod (winding_number \<gamma> z' - winding_number \<gamma> z)"
       
  4446       by simp
       
  4447     then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - \<bar>Re (winding_number \<gamma> z') - Re (winding_number \<gamma> z)\<bar>"
       
  4448       using abs_Re_le_cmod [of "winding_number \<gamma> z' - winding_number \<gamma> z"] by simp
       
  4449     then have wnz_12': "\<bar>Re (winding_number \<gamma> z')\<bar> > 1/2"
       
  4450       by linarith
       
  4451     moreover have "\<bar>Re (winding_number \<gamma> z')\<bar> < 1/2"
       
  4452       apply (rule winding_number_lt_half [OF \<gamma> *])
       
  4453       using azb \<open>d>0\<close> pag
       
  4454       apply (auto simp: add_strict_increasing anz divide_simps algebra_simps dest!: subsetD)
       
  4455       done
       
  4456     ultimately have False
       
  4457       by simp
       
  4458   }
       
  4459   then show ?thesis by force
       
  4460 qed
       
  4461 
       
  4462 lemma winding_number_lt_half_linepath: "z \<notin> closed_segment a b \<Longrightarrow> \<bar>Re (winding_number (linepath a b) z)\<bar> < 1/2"
       
  4463   using separating_hyperplane_closed_point [of "closed_segment a b" z]
       
  4464   apply auto
       
  4465   apply (simp add: closed_segment_def)
       
  4466   apply (drule less_imp_le)
       
  4467   apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]])
       
  4468   apply (auto simp: segment)
       
  4469   done
       
  4470 
       
  4471 
       
  4472 text\<open> Positivity of WN for a linepath.\<close>
       
  4473 lemma winding_number_linepath_pos_lt:
       
  4474     assumes "0 < Im ((b - a) * cnj (b - z))"
       
  4475       shows "0 < Re(winding_number(linepath a b) z)"
       
  4476 proof -
       
  4477   have z: "z \<notin> path_image (linepath a b)"
       
  4478     using assms
       
  4479     by (simp add: closed_segment_def) (force simp: algebra_simps)
       
  4480   show ?thesis
       
  4481     apply (rule winding_number_pos_lt [OF valid_path_linepath z assms])
       
  4482     apply (simp add: linepath_def algebra_simps)
       
  4483     done
       
  4484 qed
       
  4485 
       
  4486 
       
  4487 subsection\<open>Cauchy's integral formula, again for a convex enclosing set.\<close>
       
  4488 
       
  4489 lemma Cauchy_integral_formula_weak:
       
  4490     assumes s: "convex s" and "finite k" and conf: "continuous_on s f"
       
  4491         and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)"
       
  4492         and z: "z \<in> interior s - k" and vpg: "valid_path \<gamma>"
       
  4493         and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  4494       shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
       
  4495 proof -
       
  4496   obtain f' where f': "(f has_field_derivative f') (at z)"
       
  4497     using fcd [OF z] by (auto simp: field_differentiable_def)
       
  4498   have pas: "path_image \<gamma> \<subseteq> s" and znotin: "z \<notin> path_image \<gamma>" using pasz by blast+
       
  4499   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
       
  4500   proof (cases "x = z")
       
  4501     case True then show ?thesis
       
  4502       apply (simp add: continuous_within)
       
  4503       apply (rule Lim_transform_away_within [of _ "z+1" _ "\<lambda>w::complex. (f w - f z)/(w - z)"])
       
  4504       using has_field_derivative_at_within DERIV_within_iff f'
       
  4505       apply (fastforce simp add:)+
       
  4506       done
       
  4507   next
       
  4508     case False
       
  4509     then have dxz: "dist x z > 0" by auto
       
  4510     have cf: "continuous (at x within s) f"
       
  4511       using conf continuous_on_eq_continuous_within that by blast
       
  4512     have "continuous (at x within s) (\<lambda>w. (f w - f z) / (w - z))"
       
  4513       by (rule cf continuous_intros | simp add: False)+
       
  4514     then show ?thesis
       
  4515       apply (rule continuous_transform_within [OF _ dxz that, of "\<lambda>w::complex. (f w - f z)/(w - z)"])
       
  4516       apply (force simp: dist_commute)
       
  4517       done
       
  4518   qed
       
  4519   have fink': "finite (insert z k)" using \<open>finite k\<close> by blast
       
  4520   have *: "((\<lambda>w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>"
       
  4521     apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop])
       
  4522     using c apply (force simp: continuous_on_eq_continuous_within)
       
  4523     apply (rename_tac w)
       
  4524     apply (rule_tac d="dist w z" and f = "\<lambda>w. (f w - f z)/(w - z)" in field_differentiable_transform_within)
       
  4525     apply (simp_all add: dist_pos_lt dist_commute)
       
  4526     apply (metis less_irrefl)
       
  4527     apply (rule derivative_intros fcd | simp)+
       
  4528     done
       
  4529   show ?thesis
       
  4530     apply (rule has_contour_integral_eq)
       
  4531     using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *]
       
  4532     apply (auto simp: mult_ac divide_simps)
       
  4533     done
       
  4534 qed
       
  4535 
       
  4536 theorem Cauchy_integral_formula_convex_simple:
       
  4537     "\<lbrakk>convex s; f holomorphic_on s; z \<in> interior s; valid_path \<gamma>; path_image \<gamma> \<subseteq> s - {z};
       
  4538       pathfinish \<gamma> = pathstart \<gamma>\<rbrakk>
       
  4539      \<Longrightarrow> ((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
       
  4540   apply (rule Cauchy_integral_formula_weak [where k = "{}"])
       
  4541   using holomorphic_on_imp_continuous_on
       
  4542   by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE)
       
  4543 
       
  4544 
       
  4545 subsection\<open>Homotopy forms of Cauchy's theorem\<close>
       
  4546 
       
  4547 proposition Cauchy_theorem_homotopic:
       
  4548     assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h"
       
  4549         and "open s" and f: "f holomorphic_on s"
       
  4550         and vpg: "valid_path g" and vph: "valid_path h"
       
  4551     shows "contour_integral g f = contour_integral h f"
       
  4552 proof -
       
  4553   have pathsf: "linked_paths atends g h"
       
  4554     using hom  by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop)
       
  4555   obtain k :: "real \<times> real \<Rightarrow> complex"
       
  4556     where contk: "continuous_on ({0..1} \<times> {0..1}) k"
       
  4557       and ks: "k ` ({0..1} \<times> {0..1}) \<subseteq> s"
       
  4558       and k [simp]: "\<forall>x. k (0, x) = g x" "\<forall>x. k (1, x) = h x"
       
  4559       and ksf: "\<forall>t\<in>{0..1}. linked_paths atends g (\<lambda>x. k (t, x))"
       
  4560       using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm)
       
  4561   have ucontk: "uniformly_continuous_on ({0..1} \<times> {0..1}) k"
       
  4562     by (blast intro: compact_Times compact_uniformly_continuous [OF contk])
       
  4563   { fix t::real assume t: "t \<in> {0..1}"
       
  4564     have pak: "path (k o (\<lambda>u. (t, u)))"
       
  4565       unfolding path_def
       
  4566       apply (rule continuous_intros continuous_on_subset [OF contk])+
       
  4567       using t by force
       
  4568     have pik: "path_image (k \<circ> Pair t) \<subseteq> s"
       
  4569       using ks t by (auto simp: path_image_def)
       
  4570     obtain e where "e>0" and e:
       
  4571          "\<And>g h. \<lbrakk>valid_path g; valid_path h;
       
  4572                   \<forall>u\<in>{0..1}. cmod (g u - (k \<circ> Pair t) u) < e \<and> cmod (h u - (k \<circ> Pair t) u) < e;
       
  4573                   linked_paths atends g h\<rbrakk>
       
  4574                  \<Longrightarrow> contour_integral h f = contour_integral g f"
       
  4575       using contour_integral_nearby [OF \<open>open s\<close> pak pik, of atends] f by metis
       
  4576     obtain d where "d>0" and d:
       
  4577         "\<And>x x'. \<lbrakk>x \<in> {0..1} \<times> {0..1}; x' \<in> {0..1} \<times> {0..1}; norm (x'-x) < d\<rbrakk> \<Longrightarrow> norm (k x' - k x) < e/4"
       
  4578       by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \<open>e>0\<close>)
       
  4579     { fix t1 t2
       
  4580       assume t1: "0 \<le> t1" "t1 \<le> 1" and t2: "0 \<le> t2" "t2 \<le> 1" and ltd: "\<bar>t1 - t\<bar> < d" "\<bar>t2 - t\<bar> < d"
       
  4581       have no2: "\<And>g1 k1 kt. \<lbrakk>norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\<rbrakk> \<Longrightarrow> norm(g1 - kt) < e"
       
  4582         using \<open>e > 0\<close>
       
  4583         apply (rule_tac y = k1 in norm_triangle_half_l)
       
  4584         apply (auto simp: norm_minus_commute intro: order_less_trans)
       
  4585         done
       
  4586       have "\<exists>d>0. \<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
       
  4587                           (\<forall>u\<in>{0..1}. cmod (g1 u - k (t1, u)) < d \<and> cmod (g2 u - k (t2, u)) < d) \<and>
       
  4588                           linked_paths atends g1 g2 \<longrightarrow>
       
  4589                           contour_integral g2 f = contour_integral g1 f"
       
  4590         apply (rule_tac x="e/4" in exI)
       
  4591         using t t1 t2 ltd \<open>e > 0\<close>
       
  4592         apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1)
       
  4593         done
       
  4594     }
       
  4595     then have "\<exists>e. 0 < e \<and>
       
  4596               (\<forall>t1 t2. t1 \<in> {0..1} \<and> t2 \<in> {0..1} \<and> \<bar>t1 - t\<bar> < e \<and> \<bar>t2 - t\<bar> < e
       
  4597                 \<longrightarrow> (\<exists>d. 0 < d \<and>
       
  4598                      (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
       
  4599                        (\<forall>u \<in> {0..1}.
       
  4600                           norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and>
       
  4601                           linked_paths atends g1 g2
       
  4602                           \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
       
  4603       by (rule_tac x=d in exI) (simp add: \<open>d > 0\<close>)
       
  4604   }
       
  4605   then obtain ee where ee:
       
  4606        "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0 \<and>
       
  4607           (\<forall>t1 t2. t1 \<in> {0..1} \<longrightarrow> t2 \<in> {0..1} \<longrightarrow> \<bar>t1 - t\<bar> < ee t \<longrightarrow> \<bar>t2 - t\<bar> < ee t
       
  4608             \<longrightarrow> (\<exists>d. 0 < d \<and>
       
  4609                  (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
       
  4610                    (\<forall>u \<in> {0..1}.
       
  4611                       norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and>
       
  4612                       linked_paths atends g1 g2
       
  4613                       \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
       
  4614     by metis
       
  4615   note ee_rule = ee [THEN conjunct2, rule_format]
       
  4616   define C where "C = (\<lambda>t. ball t (ee t / 3)) ` {0..1}"
       
  4617   have "\<forall>t \<in> C. open t" by (simp add: C_def)
       
  4618   moreover have "{0..1} \<subseteq> \<Union>C"
       
  4619     using ee [THEN conjunct1] by (auto simp: C_def dist_norm)
       
  4620   ultimately obtain C' where C': "C' \<subseteq> C" "finite C'" and C'01: "{0..1} \<subseteq> \<Union>C'"
       
  4621     by (rule compactE [OF compact_interval])
       
  4622   define kk where "kk = {t \<in> {0..1}. ball t (ee t / 3) \<in> C'}"
       
  4623   have kk01: "kk \<subseteq> {0..1}" by (auto simp: kk_def)
       
  4624   define e where "e = Min (ee ` kk)"
       
  4625   have C'_eq: "C' = (\<lambda>t. ball t (ee t / 3)) ` kk"
       
  4626     using C' by (auto simp: kk_def C_def)
       
  4627   have ee_pos[simp]: "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0"
       
  4628     by (simp add: kk_def ee)
       
  4629   moreover have "finite kk"
       
  4630     using \<open>finite C'\<close> kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD)
       
  4631   moreover have "kk \<noteq> {}" using \<open>{0..1} \<subseteq> \<Union>C'\<close> C'_eq by force
       
  4632   ultimately have "e > 0"
       
  4633     using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def)
       
  4634   then obtain N::nat where "N > 0" and N: "1/N < e/3"
       
  4635     by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral)
       
  4636   have e_le_ee: "\<And>i. i \<in> kk \<Longrightarrow> e \<le> ee i"
       
  4637     using \<open>finite kk\<close> by (simp add: e_def Min_le_iff [of "ee ` kk"])
       
  4638   have plus: "\<exists>t \<in> kk. x \<in> ball t (ee t / 3)" if "x \<in> {0..1}" for x
       
  4639     using C' subsetD [OF C'01 that]  unfolding C'_eq by blast
       
  4640   have [OF order_refl]:
       
  4641       "\<exists>d. 0 < d \<and> (\<forall>j. valid_path j \<and> (\<forall>u \<in> {0..1}. norm(j u - k (n/N, u)) < d) \<and> linked_paths atends g j
       
  4642                         \<longrightarrow> contour_integral j f = contour_integral g f)"
       
  4643        if "n \<le> N" for n
       
  4644   using that
       
  4645   proof (induct n)
       
  4646     case 0 show ?case using ee_rule [of 0 0 0]
       
  4647       apply clarsimp
       
  4648       apply (rule_tac x=d in exI, safe)
       
  4649       by (metis diff_self vpg norm_zero)
       
  4650   next
       
  4651     case (Suc n)
       
  4652     then have N01: "n/N \<in> {0..1}" "(Suc n)/N \<in> {0..1}"  by auto
       
  4653     then obtain t where t: "t \<in> kk" "n/N \<in> ball t (ee t / 3)"
       
  4654       using plus [of "n/N"] by blast
       
  4655     then have nN_less: "\<bar>n/N - t\<bar> < ee t"
       
  4656       by (simp add: dist_norm del: less_divide_eq_numeral1)
       
  4657     have n'N_less: "\<bar>real (Suc n) / real N - t\<bar> < ee t"
       
  4658       using t N \<open>N > 0\<close> e_le_ee [of t]
       
  4659       by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps)
       
  4660     have t01: "t \<in> {0..1}" using \<open>kk \<subseteq> {0..1}\<close> \<open>t \<in> kk\<close> by blast
       
  4661     obtain d1 where "d1 > 0" and d1:
       
  4662         "\<And>g1 g2. \<lbrakk>valid_path g1; valid_path g2;
       
  4663                    \<forall>u\<in>{0..1}. cmod (g1 u - k (n/N, u)) < d1 \<and> cmod (g2 u - k ((Suc n) / N, u)) < d1;
       
  4664                    linked_paths atends g1 g2\<rbrakk>
       
  4665                    \<Longrightarrow> contour_integral g2 f = contour_integral g1 f"
       
  4666       using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce
       
  4667     have "n \<le> N" using Suc.prems by auto
       
  4668     with Suc.hyps
       
  4669     obtain d2 where "d2 > 0"
       
  4670       and d2: "\<And>j. \<lbrakk>valid_path j; \<forall>u\<in>{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\<rbrakk>
       
  4671                      \<Longrightarrow> contour_integral j f = contour_integral g f"
       
  4672         by auto
       
  4673     have "continuous_on {0..1} (k o (\<lambda>u. (n/N, u)))"
       
  4674       apply (rule continuous_intros continuous_on_subset [OF contk])+
       
  4675       using N01 by auto
       
  4676     then have pkn: "path (\<lambda>u. k (n/N, u))"
       
  4677       by (simp add: path_def)
       
  4678     have min12: "min d1 d2 > 0" by (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>)
       
  4679     obtain p where "polynomial_function p"
       
  4680         and psf: "pathstart p = pathstart (\<lambda>u. k (n/N, u))"
       
  4681                  "pathfinish p = pathfinish (\<lambda>u. k (n/N, u))"
       
  4682         and pk_le:  "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (p t - k (n/N, t)) < min d1 d2"
       
  4683       using path_approx_polynomial_function [OF pkn min12] by blast
       
  4684     then have vpp: "valid_path p" using valid_path_polynomial_function by blast
       
  4685     have lpa: "linked_paths atends g p"
       
  4686       by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf)
       
  4687     show ?case
       
  4688       apply (rule_tac x="min d1 d2" in exI)
       
  4689       apply (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>, clarify)
       
  4690       apply (rule_tac s="contour_integral p f" in trans)
       
  4691       using pk_le N01(1) ksf pathfinish_def pathstart_def
       
  4692       apply (force intro!: vpp d1 simp add: linked_paths_def psf ksf)
       
  4693       using pk_le N01 apply (force intro!: vpp d2 lpa simp add: linked_paths_def psf ksf)
       
  4694       done
       
  4695   qed
       
  4696   then obtain d where "0 < d"
       
  4697                        "\<And>j. valid_path j \<and> (\<forall>u \<in> {0..1}. norm(j u - k (1,u)) < d) \<and>
       
  4698                             linked_paths atends g j
       
  4699                             \<Longrightarrow> contour_integral j f = contour_integral g f"
       
  4700     using \<open>N>0\<close> by auto
       
  4701   then have "linked_paths atends g h \<Longrightarrow> contour_integral h f = contour_integral g f"
       
  4702     using \<open>N>0\<close> vph by fastforce
       
  4703   then show ?thesis
       
  4704     by (simp add: pathsf)
       
  4705 qed
       
  4706 
       
  4707 proposition Cauchy_theorem_homotopic_paths:
       
  4708     assumes hom: "homotopic_paths s g h"
       
  4709         and "open s" and f: "f holomorphic_on s"
       
  4710         and vpg: "valid_path g" and vph: "valid_path h"
       
  4711     shows "contour_integral g f = contour_integral h f"
       
  4712   using Cauchy_theorem_homotopic [of True s g h] assms by simp
       
  4713 
       
  4714 proposition Cauchy_theorem_homotopic_loops:
       
  4715     assumes hom: "homotopic_loops s g h"
       
  4716         and "open s" and f: "f holomorphic_on s"
       
  4717         and vpg: "valid_path g" and vph: "valid_path h"
       
  4718     shows "contour_integral g f = contour_integral h f"
       
  4719   using Cauchy_theorem_homotopic [of False s g h] assms by simp
       
  4720 
       
  4721 lemma has_contour_integral_newpath:
       
  4722     "\<lbrakk>(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\<rbrakk>
       
  4723      \<Longrightarrow> (f has_contour_integral y) g"
       
  4724   using has_contour_integral_integral contour_integral_unique by auto
       
  4725 
       
  4726 lemma Cauchy_theorem_null_homotopic:
       
  4727      "\<lbrakk>f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
       
  4728   apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp)
       
  4729   using contour_integrable_holomorphic_simple
       
  4730     apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset)
       
  4731   by (simp add: Cauchy_theorem_homotopic_loops)
       
  4732 
       
  4733 
       
  4734 
       
  4735 subsection\<open>More winding number properties\<close>
       
  4736 
       
  4737 text\<open>including the fact that it's +-1 inside a simple closed curve.\<close>
       
  4738 
       
  4739 lemma winding_number_homotopic_paths:
       
  4740     assumes "homotopic_paths (-{z}) g h"
       
  4741       shows "winding_number g z = winding_number h z"
       
  4742 proof -
       
  4743   have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto
       
  4744   moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h"
       
  4745     using homotopic_paths_imp_subset [OF assms] by auto
       
  4746   ultimately obtain d e where "d > 0" "e > 0"
       
  4747       and d: "\<And>p. \<lbrakk>path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \<forall>t\<in>{0..1}. norm (p t - g t) < d\<rbrakk>
       
  4748             \<Longrightarrow> homotopic_paths (-{z}) g p"
       
  4749       and e: "\<And>q. \<lbrakk>path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \<forall>t\<in>{0..1}. norm (q t - h t) < e\<rbrakk>
       
  4750             \<Longrightarrow> homotopic_paths (-{z}) h q"
       
  4751     using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force
       
  4752   obtain p where p:
       
  4753        "valid_path p" "z \<notin> path_image p"
       
  4754        "pathstart p = pathstart g" "pathfinish p = pathfinish g"
       
  4755        and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d"
       
  4756        and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z"
       
  4757     using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast
       
  4758   obtain q where q:
       
  4759        "valid_path q" "z \<notin> path_image q"
       
  4760        "pathstart q = pathstart h" "pathfinish q = pathfinish h"
       
  4761        and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e"
       
  4762        and paq:  "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z"
       
  4763     using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast
       
  4764   have gp: "homotopic_paths (- {z}) g p"
       
  4765     by (simp add: d p valid_path_imp_path norm_minus_commute gp_less)
       
  4766   have hq: "homotopic_paths (- {z}) h q"
       
  4767     by (simp add: e q valid_path_imp_path norm_minus_commute hq_less)
       
  4768   have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))"
       
  4769     apply (rule Cauchy_theorem_homotopic_paths [of "-{z}"])
       
  4770     apply (blast intro: homotopic_paths_trans homotopic_paths_sym gp hq assms)
       
  4771     apply (auto intro!: holomorphic_intros simp: p q)
       
  4772     done
       
  4773   then show ?thesis
       
  4774     by (simp add: pap paq)
       
  4775 qed
       
  4776 
       
  4777 lemma winding_number_homotopic_loops:
       
  4778     assumes "homotopic_loops (-{z}) g h"
       
  4779       shows "winding_number g z = winding_number h z"
       
  4780 proof -
       
  4781   have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto
       
  4782   moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h"
       
  4783     using homotopic_loops_imp_subset [OF assms] by auto
       
  4784   moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h"
       
  4785     using homotopic_loops_imp_loop [OF assms] by auto
       
  4786   ultimately obtain d e where "d > 0" "e > 0"
       
  4787       and d: "\<And>p. \<lbrakk>path p; pathfinish p = pathstart p; \<forall>t\<in>{0..1}. norm (p t - g t) < d\<rbrakk>
       
  4788             \<Longrightarrow> homotopic_loops (-{z}) g p"
       
  4789       and e: "\<And>q. \<lbrakk>path q; pathfinish q = pathstart q; \<forall>t\<in>{0..1}. norm (q t - h t) < e\<rbrakk>
       
  4790             \<Longrightarrow> homotopic_loops (-{z}) h q"
       
  4791     using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force
       
  4792   obtain p where p:
       
  4793        "valid_path p" "z \<notin> path_image p"
       
  4794        "pathstart p = pathstart g" "pathfinish p = pathfinish g"
       
  4795        and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d"
       
  4796        and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z"
       
  4797     using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast
       
  4798   obtain q where q:
       
  4799        "valid_path q" "z \<notin> path_image q"
       
  4800        "pathstart q = pathstart h" "pathfinish q = pathfinish h"
       
  4801        and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e"
       
  4802        and paq:  "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z"
       
  4803     using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast
       
  4804   have gp: "homotopic_loops (- {z}) g p"
       
  4805     by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path)
       
  4806   have hq: "homotopic_loops (- {z}) h q"
       
  4807     by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path)
       
  4808   have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))"
       
  4809     apply (rule Cauchy_theorem_homotopic_loops [of "-{z}"])
       
  4810     apply (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms)
       
  4811     apply (auto intro!: holomorphic_intros simp: p q)
       
  4812     done
       
  4813   then show ?thesis
       
  4814     by (simp add: pap paq)
       
  4815 qed
       
  4816 
       
  4817 lemma winding_number_paths_linear_eq:
       
  4818   "\<lbrakk>path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g;
       
  4819     \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk>
       
  4820         \<Longrightarrow> winding_number h z = winding_number g z"
       
  4821   by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths elim: )
       
  4822 
       
  4823 lemma winding_number_loops_linear_eq:
       
  4824   "\<lbrakk>path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h;
       
  4825     \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk>
       
  4826         \<Longrightarrow> winding_number h z = winding_number g z"
       
  4827   by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops elim: )
       
  4828 
       
  4829 lemma winding_number_nearby_paths_eq:
       
  4830      "\<lbrakk>path g; path h;
       
  4831       pathstart h = pathstart g; pathfinish h = pathfinish g;
       
  4832       \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk>
       
  4833       \<Longrightarrow> winding_number h z = winding_number g z"
       
  4834   by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq)
       
  4835 
       
  4836 lemma winding_number_nearby_loops_eq:
       
  4837      "\<lbrakk>path g; path h;
       
  4838       pathfinish g = pathstart g;
       
  4839         pathfinish h = pathstart h;
       
  4840       \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk>
       
  4841       \<Longrightarrow> winding_number h z = winding_number g z"
       
  4842   by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq)
       
  4843 
       
  4844 
       
  4845 proposition winding_number_subpath_combine:
       
  4846     "\<lbrakk>path g; z \<notin> path_image g;
       
  4847       u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk>
       
  4848       \<Longrightarrow> winding_number (subpath u v g) z + winding_number (subpath v w g) z =
       
  4849           winding_number (subpath u w g) z"
       
  4850 apply (rule trans [OF winding_number_join [THEN sym]
       
  4851                       winding_number_homotopic_paths [OF homotopic_join_subpaths]])
       
  4852 apply (auto dest: path_image_subpath_subset)
       
  4853 done
       
  4854 
       
  4855 
       
  4856 subsection\<open>Partial circle path\<close>
       
  4857 
       
  4858 definition part_circlepath :: "[complex, real, real, real, real] \<Rightarrow> complex"
       
  4859   where "part_circlepath z r s t \<equiv> \<lambda>x. z + of_real r * exp (\<i> * of_real (linepath s t x))"
       
  4860 
       
  4861 lemma pathstart_part_circlepath [simp]:
       
  4862      "pathstart(part_circlepath z r s t) = z + r*exp(\<i> * s)"
       
  4863 by (metis part_circlepath_def pathstart_def pathstart_linepath)
       
  4864 
       
  4865 lemma pathfinish_part_circlepath [simp]:
       
  4866      "pathfinish(part_circlepath z r s t) = z + r*exp(\<i>*t)"
       
  4867 by (metis part_circlepath_def pathfinish_def pathfinish_linepath)
       
  4868 
       
  4869 proposition has_vector_derivative_part_circlepath [derivative_intros]:
       
  4870     "((part_circlepath z r s t) has_vector_derivative
       
  4871       (\<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)))
       
  4872      (at x within X)"
       
  4873   apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real)
       
  4874   apply (rule has_vector_derivative_real_complex)
       
  4875   apply (rule derivative_eq_intros | simp)+
       
  4876   done
       
  4877 
       
  4878 corollary vector_derivative_part_circlepath:
       
  4879     "vector_derivative (part_circlepath z r s t) (at x) =
       
  4880        \<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)"
       
  4881   using has_vector_derivative_part_circlepath vector_derivative_at by blast
       
  4882 
       
  4883 corollary vector_derivative_part_circlepath01:
       
  4884     "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
       
  4885      \<Longrightarrow> vector_derivative (part_circlepath z r s t) (at x within {0..1}) =
       
  4886           \<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)"
       
  4887   using has_vector_derivative_part_circlepath
       
  4888   by (auto simp: vector_derivative_at_within_ivl)
       
  4889 
       
  4890 lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)"
       
  4891   apply (simp add: valid_path_def)
       
  4892   apply (rule C1_differentiable_imp_piecewise)
       
  4893   apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath
       
  4894               intro!: continuous_intros)
       
  4895   done
       
  4896 
       
  4897 lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)"
       
  4898   by (simp add: valid_path_imp_path)
       
  4899 
       
  4900 proposition path_image_part_circlepath:
       
  4901   assumes "s \<le> t"
       
  4902     shows "path_image (part_circlepath z r s t) = {z + r * exp(\<i> * of_real x) | x. s \<le> x \<and> x \<le> t}"
       
  4903 proof -
       
  4904   { fix z::real
       
  4905     assume "0 \<le> z" "z \<le> 1"
       
  4906     with \<open>s \<le> t\<close> have "\<exists>x. (exp (\<i> * linepath s t z) = exp (\<i> * of_real x)) \<and> s \<le> x \<and> x \<le> t"
       
  4907       apply (rule_tac x="(1 - z) * s + z * t" in exI)
       
  4908       apply (simp add: linepath_def scaleR_conv_of_real algebra_simps)
       
  4909       apply (rule conjI)
       
  4910       using mult_right_mono apply blast
       
  4911       using affine_ineq  by (metis "mult.commute")
       
  4912   }
       
  4913   moreover
       
  4914   { fix z
       
  4915     assume "s \<le> z" "z \<le> t"
       
  4916     then have "z + of_real r * exp (\<i> * of_real z) \<in> (\<lambda>x. z + of_real r * exp (\<i> * linepath s t x)) ` {0..1}"
       
  4917       apply (rule_tac x="(z - s)/(t - s)" in image_eqI)
       
  4918       apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq)
       
  4919       apply (auto simp: algebra_simps divide_simps)
       
  4920       done
       
  4921   }
       
  4922   ultimately show ?thesis
       
  4923     by (fastforce simp add: path_image_def part_circlepath_def)
       
  4924 qed
       
  4925 
       
  4926 corollary path_image_part_circlepath_subset:
       
  4927     "\<lbrakk>s \<le> t; 0 \<le> r\<rbrakk> \<Longrightarrow> path_image(part_circlepath z r s t) \<subseteq> sphere z r"
       
  4928 by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult)
       
  4929 
       
  4930 proposition in_path_image_part_circlepath:
       
  4931   assumes "w \<in> path_image(part_circlepath z r s t)" "s \<le> t" "0 \<le> r"
       
  4932     shows "norm(w - z) = r"
       
  4933 proof -
       
  4934   have "w \<in> {c. dist z c = r}"
       
  4935     by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms)
       
  4936   thus ?thesis
       
  4937     by (simp add: dist_norm norm_minus_commute)
       
  4938 qed
       
  4939 
       
  4940 proposition finite_bounded_log: "finite {z::complex. norm z \<le> b \<and> exp z = w}"
       
  4941 proof (cases "w = 0")
       
  4942   case True then show ?thesis by auto
       
  4943 next
       
  4944   case False
       
  4945   have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \<i>) \<le> b + cmod (Ln w)}"
       
  4946     apply (simp add: norm_mult finite_int_iff_bounded_le)
       
  4947     apply (rule_tac x="\<lfloor>(b + cmod (Ln w)) / (2*pi)\<rfloor>" in exI)
       
  4948     apply (auto simp: divide_simps le_floor_iff)
       
  4949     done
       
  4950   have [simp]: "\<And>P f. {z. P z \<and> (\<exists>n. z = f n)} = f ` {n. P (f n)}"
       
  4951     by blast
       
  4952   show ?thesis
       
  4953     apply (subst exp_Ln [OF False, symmetric])
       
  4954     apply (simp add: exp_eq)
       
  4955     using norm_add_leD apply (fastforce intro: finite_subset [OF _ *])
       
  4956     done
       
  4957 qed
       
  4958 
       
  4959 lemma finite_bounded_log2:
       
  4960   fixes a::complex
       
  4961     assumes "a \<noteq> 0"
       
  4962     shows "finite {z. norm z \<le> b \<and> exp(a*z) = w}"
       
  4963 proof -
       
  4964   have *: "finite ((\<lambda>z. z / a) ` {z. cmod z \<le> b * cmod a \<and> exp z = w})"
       
  4965     by (rule finite_imageI [OF finite_bounded_log])
       
  4966   show ?thesis
       
  4967     by (rule finite_subset [OF _ *]) (force simp: assms norm_mult)
       
  4968 qed
       
  4969 
       
  4970 proposition has_contour_integral_bound_part_circlepath_strong:
       
  4971   assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)"
       
  4972       and "finite k" and le: "0 \<le> B" "0 < r" "s \<le> t"
       
  4973       and B: "\<And>x. x \<in> path_image(part_circlepath z r s t) - k \<Longrightarrow> norm(f x) \<le> B"
       
  4974     shows "cmod i \<le> B * r * (t - s)"
       
  4975 proof -
       
  4976   consider "s = t" | "s < t" using \<open>s \<le> t\<close> by linarith
       
  4977   then show ?thesis
       
  4978   proof cases
       
  4979     case 1 with fi [unfolded has_contour_integral]
       
  4980     have "i = 0"  by (simp add: vector_derivative_part_circlepath)
       
  4981     with assms show ?thesis by simp
       
  4982   next
       
  4983     case 2
       
  4984     have [simp]: "\<bar>r\<bar> = r" using \<open>r > 0\<close> by linarith
       
  4985     have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s"
       
  4986       by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff)
       
  4987     have "finite (part_circlepath z r s t -` {y} \<inter> {0..1})" if "y \<in> k" for y
       
  4988     proof -
       
  4989       define w where "w = (y - z)/of_real r / exp(\<i> * of_real s)"
       
  4990       have fin: "finite (of_real -` {z. cmod z \<le> 1 \<and> exp (\<i> * complex_of_real (t - s) * z) = w})"
       
  4991         apply (rule finite_vimageI [OF finite_bounded_log2])
       
  4992         using \<open>s < t\<close> apply (auto simp: inj_of_real)
       
  4993         done
       
  4994       show ?thesis
       
  4995         apply (simp add: part_circlepath_def linepath_def vimage_def)
       
  4996         apply (rule finite_subset [OF _ fin])
       
  4997         using le
       
  4998         apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff)
       
  4999         done
       
  5000     qed
       
  5001     then have fin01: "finite ((part_circlepath z r s t) -` k \<inter> {0..1})"
       
  5002       by (rule finite_finite_vimage_IntI [OF \<open>finite k\<close>])
       
  5003     have **: "((\<lambda>x. if (part_circlepath z r s t x) \<in> k then 0
       
  5004                     else f(part_circlepath z r s t x) *
       
  5005                        vector_derivative (part_circlepath z r s t) (at x)) has_integral i)  {0..1}"
       
  5006       apply (rule has_integral_spike
       
  5007               [where f = "\<lambda>x. f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)"])
       
  5008       apply (rule negligible_finite [OF fin01])
       
  5009       using fi has_contour_integral
       
  5010       apply auto
       
  5011       done
       
  5012     have *: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1; part_circlepath z r s t x \<notin> k\<rbrakk> \<Longrightarrow> cmod (f (part_circlepath z r s t x)) \<le> B"
       
  5013       by (auto intro!: B [unfolded path_image_def image_def, simplified])
       
  5014     show ?thesis
       
  5015       apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified])
       
  5016       using assms apply force
       
  5017       apply (simp add: norm_mult vector_derivative_part_circlepath)
       
  5018       using le * "2" \<open>r > 0\<close> by auto
       
  5019   qed
       
  5020 qed
       
  5021 
       
  5022 corollary has_contour_integral_bound_part_circlepath:
       
  5023       "\<lbrakk>(f has_contour_integral i) (part_circlepath z r s t);
       
  5024         0 \<le> B; 0 < r; s \<le> t;
       
  5025         \<And>x. x \<in> path_image(part_circlepath z r s t) \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
       
  5026        \<Longrightarrow> norm i \<le> B*r*(t - s)"
       
  5027   by (auto intro: has_contour_integral_bound_part_circlepath_strong)
       
  5028 
       
  5029 proposition contour_integrable_continuous_part_circlepath:
       
  5030      "continuous_on (path_image (part_circlepath z r s t)) f
       
  5031       \<Longrightarrow> f contour_integrable_on (part_circlepath z r s t)"
       
  5032   apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def)
       
  5033   apply (rule integrable_continuous_real)
       
  5034   apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl])
       
  5035   done
       
  5036 
       
  5037 proposition winding_number_part_circlepath_pos_less:
       
  5038   assumes "s < t" and no: "norm(w - z) < r"
       
  5039     shows "0 < Re (winding_number(part_circlepath z r s t) w)"
       
  5040 proof -
       
  5041   have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2)
       
  5042   note valid_path_part_circlepath
       
  5043   moreover have " w \<notin> path_image (part_circlepath z r s t)"
       
  5044     using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def)
       
  5045   moreover have "0 < r * (t - s) * (r - cmod (w - z))"
       
  5046     using assms by (metis \<open>0 < r\<close> diff_gt_0_iff_gt mult_pos_pos)
       
  5047   ultimately show ?thesis
       
  5048     apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"])
       
  5049     apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac)
       
  5050     apply (rule mult_left_mono)+
       
  5051     using Re_Im_le_cmod [of "w-z" "linepath s t x" for x]
       
  5052     apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square])
       
  5053     using assms \<open>0 < r\<close> by auto
       
  5054 qed
       
  5055 
       
  5056 proposition simple_path_part_circlepath:
       
  5057     "simple_path(part_circlepath z r s t) \<longleftrightarrow> (r \<noteq> 0 \<and> s \<noteq> t \<and> \<bar>s - t\<bar> \<le> 2*pi)"
       
  5058 proof (cases "r = 0 \<or> s = t")
       
  5059   case True
       
  5060   then show ?thesis
       
  5061     apply (rule disjE)
       
  5062     apply (force simp: part_circlepath_def simple_path_def intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+
       
  5063     done
       
  5064 next
       
  5065   case False then have "r \<noteq> 0" "s \<noteq> t" by auto
       
  5066   have *: "\<And>x y z s t. \<i>*((1 - x) * s + x * t) = \<i>*(((1 - y) * s + y * t)) + z  \<longleftrightarrow> \<i>*(x - y) * (t - s) = z"
       
  5067     by (simp add: algebra_simps)
       
  5068   have abs01: "\<And>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1
       
  5069                       \<Longrightarrow> (x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0 \<longleftrightarrow> \<bar>x - y\<bar> \<in> {0,1})"
       
  5070     by auto
       
  5071   have abs_away: "\<And>P. (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. P \<bar>x - y\<bar>) \<longleftrightarrow> (\<forall>x::real. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> P x)"
       
  5072     by force
       
  5073   have **: "\<And>x y. (\<exists>n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \<longleftrightarrow>
       
  5074                   (\<exists>n. \<bar>x - y\<bar> * (t - s) = 2 * (of_int n * pi))"
       
  5075     by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real]
       
  5076                     intro: exI [where x = "-n" for n])
       
  5077   have 1: "\<forall>x. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> (\<exists>n. x * (t - s) = 2 * (real_of_int n * pi)) \<longrightarrow> x = 0 \<or> x = 1 \<Longrightarrow> \<bar>s - t\<bar> \<le> 2 * pi"
       
  5078     apply (rule ccontr)
       
  5079     apply (drule_tac x="2*pi / \<bar>t - s\<bar>" in spec)
       
  5080     using False
       
  5081     apply (simp add: abs_minus_commute divide_simps)
       
  5082     apply (frule_tac x=1 in spec)
       
  5083     apply (drule_tac x="-1" in spec)
       
  5084     apply (simp add:)
       
  5085     done
       
  5086   have 2: "\<bar>s - t\<bar> = \<bar>2 * (real_of_int n * pi) / x\<bar>" if "x \<noteq> 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n
       
  5087   proof -
       
  5088     have "t-s = 2 * (real_of_int n * pi)/x"
       
  5089       using that by (simp add: field_simps)
       
  5090     then show ?thesis by (metis abs_minus_commute)
       
  5091   qed
       
  5092   show ?thesis using False
       
  5093     apply (simp add: simple_path_def path_part_circlepath)
       
  5094     apply (simp add: part_circlepath_def linepath_def exp_eq  * ** abs01  del: Set.insert_iff)
       
  5095     apply (subst abs_away)
       
  5096     apply (auto simp: 1)
       
  5097     apply (rule ccontr)
       
  5098     apply (auto simp: 2 divide_simps abs_mult dest: of_int_leD)
       
  5099     done
       
  5100 qed
       
  5101 
       
  5102 proposition arc_part_circlepath:
       
  5103   assumes "r \<noteq> 0" "s \<noteq> t" "\<bar>s - t\<bar> < 2*pi"
       
  5104     shows "arc (part_circlepath z r s t)"
       
  5105 proof -
       
  5106   have *: "x = y" if eq: "\<i> * (linepath s t x) = \<i> * (linepath s t y) + 2 * of_int n * complex_of_real pi * \<i>"
       
  5107                   and x: "x \<in> {0..1}" and y: "y \<in> {0..1}" for x y n
       
  5108     proof -
       
  5109       have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi"
       
  5110         by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_divide_cancel_left eq)
       
  5111       then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))"
       
  5112         by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re])
       
  5113       then have st: "x \<noteq> y \<Longrightarrow> (s-t) = (of_int n * (pi * 2) / (y-x))"
       
  5114         by (force simp: field_simps)
       
  5115       show ?thesis
       
  5116         apply (rule ccontr)
       
  5117         using assms x y
       
  5118         apply (simp add: st abs_mult field_simps)
       
  5119         using st
       
  5120         apply (auto simp: dest: of_int_lessD)
       
  5121         done
       
  5122     qed
       
  5123   show ?thesis
       
  5124     using assms
       
  5125     apply (simp add: arc_def)
       
  5126     apply (simp add: part_circlepath_def inj_on_def exp_eq)
       
  5127     apply (blast intro: *)
       
  5128     done
       
  5129 qed
       
  5130 
       
  5131 
       
  5132 subsection\<open>Special case of one complete circle\<close>
       
  5133 
       
  5134 definition circlepath :: "[complex, real, real] \<Rightarrow> complex"
       
  5135   where "circlepath z r \<equiv> part_circlepath z r 0 (2*pi)"
       
  5136 
       
  5137 lemma circlepath: "circlepath z r = (\<lambda>x. z + r * exp(2 * of_real pi * \<i> * of_real x))"
       
  5138   by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps)
       
  5139 
       
  5140 lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r"
       
  5141   by (simp add: circlepath_def)
       
  5142 
       
  5143 lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r"
       
  5144   by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute)
       
  5145 
       
  5146 lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)"
       
  5147 proof -
       
  5148   have "z + of_real r * exp (2 * pi * \<i> * (x + 1 / 2)) =
       
  5149         z + of_real r * exp (2 * pi * \<i> * x + pi * \<i>)"
       
  5150     by (simp add: divide_simps) (simp add: algebra_simps)
       
  5151   also have "... = z - r * exp (2 * pi * \<i> * x)"
       
  5152     by (simp add: exp_add)
       
  5153   finally show ?thesis
       
  5154     by (simp add: circlepath path_image_def sphere_def dist_norm)
       
  5155 qed
       
  5156 
       
  5157 lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x"
       
  5158   using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x]
       
  5159   by (simp add: add.commute)
       
  5160 
       
  5161 lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)"
       
  5162   using circlepath_add1 [of z r "x-1/2"]
       
  5163   by (simp add: add.commute)
       
  5164 
       
  5165 lemma path_image_circlepath_minus_subset:
       
  5166      "path_image (circlepath z (-r)) \<subseteq> path_image (circlepath z r)"
       
  5167   apply (simp add: path_image_def image_def circlepath_minus, clarify)
       
  5168   apply (case_tac "xa \<le> 1/2", force)
       
  5169   apply (force simp add: circlepath_add_half)+
       
  5170   done
       
  5171 
       
  5172 lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)"
       
  5173   using path_image_circlepath_minus_subset by fastforce
       
  5174 
       
  5175 proposition has_vector_derivative_circlepath [derivative_intros]:
       
  5176  "((circlepath z r) has_vector_derivative (2 * pi * \<i> * r * exp (2 * of_real pi * \<i> * of_real x)))
       
  5177    (at x within X)"
       
  5178   apply (simp add: circlepath_def scaleR_conv_of_real)
       
  5179   apply (rule derivative_eq_intros)
       
  5180   apply (simp add: algebra_simps)
       
  5181   done
       
  5182 
       
  5183 corollary vector_derivative_circlepath:
       
  5184    "vector_derivative (circlepath z r) (at x) =
       
  5185     2 * pi * \<i> * r * exp(2 * of_real pi * \<i> * x)"
       
  5186 using has_vector_derivative_circlepath vector_derivative_at by blast
       
  5187 
       
  5188 corollary vector_derivative_circlepath01:
       
  5189     "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
       
  5190      \<Longrightarrow> vector_derivative (circlepath z r) (at x within {0..1}) =
       
  5191           2 * pi * \<i> * r * exp(2 * of_real pi * \<i> * x)"
       
  5192   using has_vector_derivative_circlepath
       
  5193   by (auto simp: vector_derivative_at_within_ivl)
       
  5194 
       
  5195 lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)"
       
  5196   by (simp add: circlepath_def)
       
  5197 
       
  5198 lemma path_circlepath [simp]: "path (circlepath z r)"
       
  5199   by (simp add: valid_path_imp_path)
       
  5200 
       
  5201 lemma path_image_circlepath_nonneg:
       
  5202   assumes "0 \<le> r" shows "path_image (circlepath z r) = sphere z r"
       
  5203 proof -
       
  5204   have *: "x \<in> (\<lambda>u. z + (cmod (x - z)) * exp (\<i> * (of_real u * (of_real pi * 2)))) ` {0..1}" for x
       
  5205   proof (cases "x = z")
       
  5206     case True then show ?thesis by force
       
  5207   next
       
  5208     case False
       
  5209     define w where "w = x - z"
       
  5210     then have "w \<noteq> 0" by (simp add: False)
       
  5211     have **: "\<And>t. \<lbrakk>Re w = cos t * cmod w; Im w = sin t * cmod w\<rbrakk> \<Longrightarrow> w = of_real (cmod w) * exp (\<i> * t)"
       
  5212       using cis_conv_exp complex_eq_iff by auto
       
  5213     show ?thesis
       
  5214       apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"])
       
  5215       apply (simp add: divide_simps \<open>w \<noteq> 0\<close> cmod_power2 [symmetric])
       
  5216       apply (rule_tac x="t / (2*pi)" in image_eqI)
       
  5217       apply (simp add: divide_simps \<open>w \<noteq> 0\<close>)
       
  5218       using False **
       
  5219       apply (auto simp: w_def)
       
  5220       done
       
  5221   qed
       
  5222   show ?thesis
       
  5223     unfolding circlepath path_image_def sphere_def dist_norm
       
  5224     by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *)
       
  5225 qed
       
  5226 
       
  5227 proposition path_image_circlepath [simp]:
       
  5228     "path_image (circlepath z r) = sphere z \<bar>r\<bar>"
       
  5229   using path_image_circlepath_minus
       
  5230   by (force simp add: path_image_circlepath_nonneg abs_if)
       
  5231 
       
  5232 lemma has_contour_integral_bound_circlepath_strong:
       
  5233       "\<lbrakk>(f has_contour_integral i) (circlepath z r);
       
  5234         finite k; 0 \<le> B; 0 < r;
       
  5235         \<And>x. \<lbrakk>norm(x - z) = r; x \<notin> k\<rbrakk> \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
       
  5236         \<Longrightarrow> norm i \<le> B*(2*pi*r)"
       
  5237   unfolding circlepath_def
       
  5238   by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong)
       
  5239 
       
  5240 corollary has_contour_integral_bound_circlepath:
       
  5241       "\<lbrakk>(f has_contour_integral i) (circlepath z r);
       
  5242         0 \<le> B; 0 < r; \<And>x. norm(x - z) = r \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
       
  5243         \<Longrightarrow> norm i \<le> B*(2*pi*r)"
       
  5244   by (auto intro: has_contour_integral_bound_circlepath_strong)
       
  5245 
       
  5246 proposition contour_integrable_continuous_circlepath:
       
  5247     "continuous_on (path_image (circlepath z r)) f
       
  5248      \<Longrightarrow> f contour_integrable_on (circlepath z r)"
       
  5249   by (simp add: circlepath_def contour_integrable_continuous_part_circlepath)
       
  5250 
       
  5251 lemma simple_path_circlepath: "simple_path(circlepath z r) \<longleftrightarrow> (r \<noteq> 0)"
       
  5252   by (simp add: circlepath_def simple_path_part_circlepath)
       
  5253 
       
  5254 lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \<Longrightarrow> w \<notin> path_image (circlepath z r)"
       
  5255   by (simp add: sphere_def dist_norm norm_minus_commute)
       
  5256 
       
  5257 proposition contour_integral_circlepath:
       
  5258      "0 < r \<Longrightarrow> contour_integral (circlepath z r) (\<lambda>w. 1 / (w - z)) = 2 * complex_of_real pi * \<i>"
       
  5259   apply (rule contour_integral_unique)
       
  5260   apply (simp add: has_contour_integral_def)
       
  5261   apply (subst has_integral_cong)
       
  5262   apply (simp add: vector_derivative_circlepath01)
       
  5263   using has_integral_const_real [of _ 0 1]
       
  5264   apply (force simp: circlepath)
       
  5265   done
       
  5266 
       
  5267 lemma winding_number_circlepath_centre: "0 < r \<Longrightarrow> winding_number (circlepath z r) z = 1"
       
  5268   apply (rule winding_number_unique_loop)
       
  5269   apply (simp_all add: sphere_def valid_path_imp_path)
       
  5270   apply (rule_tac x="circlepath z r" in exI)
       
  5271   apply (simp add: sphere_def contour_integral_circlepath)
       
  5272   done
       
  5273 
       
  5274 proposition winding_number_circlepath:
       
  5275   assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1"
       
  5276 proof (cases "w = z")
       
  5277   case True then show ?thesis
       
  5278     using assms winding_number_circlepath_centre by auto
       
  5279 next
       
  5280   case False
       
  5281   have [simp]: "r > 0"
       
  5282     using assms le_less_trans norm_ge_zero by blast
       
  5283   define r' where "r' = norm(w - z)"
       
  5284   have "r' < r"
       
  5285     by (simp add: assms r'_def)
       
  5286   have disjo: "cball z r' \<inter> sphere z r = {}"
       
  5287     using \<open>r' < r\<close> by (force simp: cball_def sphere_def)
       
  5288   have "winding_number(circlepath z r) w = winding_number(circlepath z r) z"
       
  5289     apply (rule winding_number_around_inside [where s = "cball z r'"])
       
  5290     apply (simp_all add: disjo order.strict_implies_order winding_number_circlepath_centre)
       
  5291     apply (simp_all add: False r'_def dist_norm norm_minus_commute)
       
  5292     done
       
  5293   also have "... = 1"
       
  5294     by (simp add: winding_number_circlepath_centre)
       
  5295   finally show ?thesis .
       
  5296 qed
       
  5297 
       
  5298 
       
  5299 text\<open> Hence the Cauchy formula for points inside a circle.\<close>
       
  5300 
       
  5301 theorem Cauchy_integral_circlepath:
       
  5302   assumes "continuous_on (cball z r) f" "f holomorphic_on (ball z r)" "norm(w - z) < r"
       
  5303   shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w))
       
  5304          (circlepath z r)"
       
  5305 proof -
       
  5306   have "r > 0"
       
  5307     using assms le_less_trans norm_ge_zero by blast
       
  5308   have "((\<lambda>u. f u / (u - w)) has_contour_integral (2 * pi) * \<i> * winding_number (circlepath z r) w * f w)
       
  5309         (circlepath z r)"
       
  5310     apply (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"])
       
  5311     using assms  \<open>r > 0\<close>
       
  5312     apply (simp_all add: dist_norm norm_minus_commute)
       
  5313     apply (metis at_within_interior dist_norm holomorphic_on_def interior_ball mem_ball norm_minus_commute)
       
  5314     apply (simp add: cball_def sphere_def dist_norm, clarify)
       
  5315     apply (simp add:)
       
  5316     by (metis dist_commute dist_norm less_irrefl)
       
  5317   then show ?thesis
       
  5318     by (simp add: winding_number_circlepath assms)
       
  5319 qed
       
  5320 
       
  5321 corollary Cauchy_integral_circlepath_simple:
       
  5322   assumes "f holomorphic_on cball z r" "norm(w - z) < r"
       
  5323   shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w))
       
  5324          (circlepath z r)"
       
  5325 using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath)
       
  5326 
       
  5327 
       
  5328 lemma no_bounded_connected_component_imp_winding_number_zero:
       
  5329   assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
       
  5330       and nb: "\<And>z. bounded (connected_component_set (- s) z) \<longrightarrow> z \<in> s"
       
  5331   shows "winding_number g z = 0"
       
  5332 apply (rule winding_number_zero_in_outside)
       
  5333 apply (simp_all add: assms)
       
  5334 by (metis nb [of z] \<open>path_image g \<subseteq> s\<close> \<open>z \<notin> s\<close> contra_subsetD mem_Collect_eq outside outside_mono)
       
  5335 
       
  5336 lemma no_bounded_path_component_imp_winding_number_zero:
       
  5337   assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
       
  5338       and nb: "\<And>z. bounded (path_component_set (- s) z) \<longrightarrow> z \<in> s"
       
  5339   shows "winding_number g z = 0"
       
  5340 apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g])
       
  5341 by (simp add: bounded_subset nb path_component_subset_connected_component)
       
  5342 
       
  5343 
       
  5344 subsection\<open> Uniform convergence of path integral\<close>
       
  5345 
       
  5346 text\<open>Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\<close>
       
  5347 
       
  5348 proposition contour_integral_uniform_limit:
       
  5349   assumes ev_fint: "eventually (\<lambda>n::'a. (f n) contour_integrable_on \<gamma>) F"
       
  5350       and ev_no: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> path_image \<gamma>. norm(f n x - l x) < e) F"
       
  5351       and noleB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
       
  5352       and \<gamma>: "valid_path \<gamma>"
       
  5353       and [simp]: "~ (trivial_limit F)"
       
  5354   shows "l contour_integrable_on \<gamma>" "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F"
       
  5355 proof -
       
  5356   have "0 \<le> B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one)
       
  5357   { fix e::real
       
  5358     assume "0 < e"
       
  5359     then have eB: "0 < e / (\<bar>B\<bar> + 1)" by simp
       
  5360     obtain a where fga: "\<And>x. x \<in> {0..1} \<Longrightarrow> cmod (f a (\<gamma> x) - l (\<gamma> x)) < e / (\<bar>B\<bar> + 1)"
       
  5361                and inta: "(\<lambda>t. f a (\<gamma> t) * vector_derivative \<gamma> (at t)) integrable_on {0..1}"
       
  5362       using eventually_happens [OF eventually_conj [OF ev_no [OF eB] ev_fint]]
       
  5363       by (fastforce simp: contour_integrable_on path_image_def)
       
  5364     have Ble: "B * e / (\<bar>B\<bar> + 1) \<le> e"
       
  5365       using \<open>0 \<le> B\<close>  \<open>0 < e\<close> by (simp add: divide_simps)
       
  5366     have "\<exists>h. (\<forall>x\<in>{0..1}. cmod (l (\<gamma> x) * vector_derivative \<gamma> (at x) - h x) \<le> e) \<and> h integrable_on {0..1}"
       
  5367       apply (rule_tac x="\<lambda>x. f (a::'a) (\<gamma> x) * vector_derivative \<gamma> (at x)" in exI)
       
  5368       apply (intro inta conjI ballI)
       
  5369       apply (rule order_trans [OF _ Ble])
       
  5370       apply (frule noleB)
       
  5371       apply (frule fga)
       
  5372       using \<open>0 \<le> B\<close>  \<open>0 < e\<close>
       
  5373       apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps)
       
  5374       apply (drule (1) mult_mono [OF less_imp_le])
       
  5375       apply (simp_all add: mult_ac)
       
  5376       done
       
  5377   }
       
  5378   then show lintg: "l contour_integrable_on \<gamma>"
       
  5379     apply (simp add: contour_integrable_on)
       
  5380     apply (blast intro: integrable_uniform_limit_real)
       
  5381     done
       
  5382   { fix e::real
       
  5383     define B' where "B' = B + 1"
       
  5384     have B': "B' > 0" "B' > B" using  \<open>0 \<le> B\<close> by (auto simp: B'_def)
       
  5385     assume "0 < e"
       
  5386     then have ev_no': "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image \<gamma>. 2 * cmod (f n x - l x) < e / B'"
       
  5387       using ev_no [of "e / B' / 2"] B' by (simp add: field_simps)
       
  5388     have ie: "integral {0..1::real} (\<lambda>x. e / 2) < e" using \<open>0 < e\<close> by simp
       
  5389     have *: "cmod (f x (\<gamma> t) * vector_derivative \<gamma> (at t) - l (\<gamma> t) * vector_derivative \<gamma> (at t)) \<le> e / 2"
       
  5390              if t: "t\<in>{0..1}" and leB': "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) < e / B'" for x t
       
  5391     proof -
       
  5392       have "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) * cmod (vector_derivative \<gamma> (at t)) \<le> e * (B/ B')"
       
  5393         using mult_mono [OF less_imp_le [OF leB'] noleB] B' \<open>0 < e\<close> t by auto
       
  5394       also have "... < e"
       
  5395         by (simp add: B' \<open>0 < e\<close> mult_imp_div_pos_less)
       
  5396       finally have "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) * cmod (vector_derivative \<gamma> (at t)) < e" .
       
  5397       then show ?thesis
       
  5398         by (simp add: left_diff_distrib [symmetric] norm_mult)
       
  5399     qed
       
  5400     have "\<forall>\<^sub>F x in F. dist (contour_integral \<gamma> (f x)) (contour_integral \<gamma> l) < e"
       
  5401       apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]])
       
  5402       apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral)
       
  5403       apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric], clarify)
       
  5404       apply (rule le_less_trans [OF integral_norm_bound_integral ie])
       
  5405       apply (simp add: lintg integrable_diff contour_integrable_on [symmetric])
       
  5406       apply (blast intro: *)+
       
  5407       done
       
  5408   }
       
  5409   then show "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F"
       
  5410     by (rule tendstoI)
       
  5411 qed
       
  5412 
       
  5413 proposition contour_integral_uniform_limit_circlepath:
       
  5414   assumes ev_fint: "eventually (\<lambda>n::'a. (f n) contour_integrable_on (circlepath z r)) F"
       
  5415       and ev_no: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> path_image (circlepath z r). norm(f n x - l x) < e) F"
       
  5416       and [simp]: "~ (trivial_limit F)" "0 < r"
       
  5417   shows "l contour_integrable_on (circlepath z r)" "((\<lambda>n. contour_integral (circlepath z r) (f n)) \<longlongrightarrow> contour_integral (circlepath z r) l) F"
       
  5418 by (auto simp: vector_derivative_circlepath norm_mult intro: contour_integral_uniform_limit assms)
       
  5419 
       
  5420 
       
  5421 subsection\<open> General stepping result for derivative formulas.\<close>
       
  5422 
       
  5423 lemma sum_sqs_eq:
       
  5424   fixes x::"'a::idom" shows "x * x + y * y = x * (y * 2) \<Longrightarrow> y = x"
       
  5425   by algebra
       
  5426 
       
  5427 proposition Cauchy_next_derivative:
       
  5428   assumes "continuous_on (path_image \<gamma>) f'"
       
  5429       and leB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
       
  5430       and int: "\<And>w. w \<in> s - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f' u / (u - w)^k) has_contour_integral f w) \<gamma>"
       
  5431       and k: "k \<noteq> 0"
       
  5432       and "open s"
       
  5433       and \<gamma>: "valid_path \<gamma>"
       
  5434       and w: "w \<in> s - path_image \<gamma>"
       
  5435     shows "(\<lambda>u. f' u / (u - w)^(Suc k)) contour_integrable_on \<gamma>"
       
  5436       and "(f has_field_derivative (k * contour_integral \<gamma> (\<lambda>u. f' u/(u - w)^(Suc k))))
       
  5437            (at w)"  (is "?thes2")
       
  5438 proof -
       
  5439   have "open (s - path_image \<gamma>)" using \<open>open s\<close> closed_valid_path_image \<gamma> by blast
       
  5440   then obtain d where "d>0" and d: "ball w d \<subseteq> s - path_image \<gamma>" using w
       
  5441     using open_contains_ball by blast
       
  5442   have [simp]: "\<And>n. cmod (1 + of_nat n) = 1 + of_nat n"
       
  5443     by (metis norm_of_nat of_nat_Suc)
       
  5444   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)
       
  5445                          contour_integrable_on \<gamma>"
       
  5446     apply (simp add: eventually_at)
       
  5447     apply (rule_tac x=d in exI)
       
  5448     apply (simp add: \<open>d > 0\<close> dist_norm field_simps, clarify)
       
  5449     apply (rule contour_integrable_div [OF contour_integrable_diff])
       
  5450     using int w d
       
  5451     apply (force simp:  dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+
       
  5452     done
       
  5453   have bim_g: "bounded (image f' (path_image \<gamma>))"
       
  5454     by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms)
       
  5455   then obtain C where "C > 0" and C: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> cmod (f' (\<gamma> x)) \<le> C"
       
  5456     by (force simp: bounded_pos path_image_def)
       
  5457   have twom: "\<forall>\<^sub>F n in at w.
       
  5458                \<forall>x\<in>path_image \<gamma>.
       
  5459                 cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e"
       
  5460          if "0 < e" for e
       
  5461   proof -
       
  5462     have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k)   < e"
       
  5463             if x: "x \<in> path_image \<gamma>" and "u \<noteq> w" and uwd: "cmod (u - w) < d/2"
       
  5464                 and uw_less: "cmod (u - w) < e * (d / 2) ^ (k+2) / (1 + real k)"
       
  5465             for u x
       
  5466     proof -
       
  5467       define ff where [abs_def]:
       
  5468         "ff n w =
       
  5469           (if n = 0 then inverse(x - w)^k
       
  5470            else if n = 1 then k / (x - w)^(Suc k)
       
  5471            else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w
       
  5472       have km1: "\<And>z::complex. z \<noteq> 0 \<Longrightarrow> z ^ (k - Suc 0) = z ^ k / z"
       
  5473         by (simp add: field_simps) (metis Suc_pred \<open>k \<noteq> 0\<close> neq0_conv power_Suc)
       
  5474       have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d / 2))"
       
  5475               if "z \<in> ball w (d / 2)" "i \<le> 1" for i z
       
  5476       proof -
       
  5477         have "z \<notin> path_image \<gamma>"
       
  5478           using \<open>x \<in> path_image \<gamma>\<close> d that ball_divide_subset_numeral by blast
       
  5479         then have xz[simp]: "x \<noteq> z" using \<open>x \<in> path_image \<gamma>\<close> by blast
       
  5480         then have neq: "x * x + z * z \<noteq> x * (z * 2)"
       
  5481           by (blast intro: dest!: sum_sqs_eq)
       
  5482         with xz have "\<And>v. v \<noteq> 0 \<Longrightarrow> (x * x + z * z) * v \<noteq> (x * (z * 2) * v)" by auto
       
  5483         then have neqq: "\<And>v. v \<noteq> 0 \<Longrightarrow> x * (x * v) + z * (z * v) \<noteq> x * (z * (2 * v))"
       
  5484           by (simp add: algebra_simps)
       
  5485         show ?thesis using \<open>i \<le> 1\<close>
       
  5486           apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe)
       
  5487           apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+
       
  5488           done
       
  5489       qed
       
  5490       { fix a::real and b::real assume ab: "a > 0" "b > 0"
       
  5491         then have "k * (1 + real k) * (1 / a) \<le> k * (1 + real k) * (4 / b) \<longleftrightarrow> b \<le> 4 * a"
       
  5492           apply (subst mult_le_cancel_left_pos)
       
  5493           using \<open>k \<noteq> 0\<close>
       
  5494           apply (auto simp: divide_simps)
       
  5495           done
       
  5496         with ab have "real k * (1 + real k) / a \<le> (real k * 4 + real k * real k * 4) / b \<longleftrightarrow> b \<le> 4 * a"
       
  5497           by (simp add: field_simps)
       
  5498       } note canc = this
       
  5499       have ff2: "cmod (ff (Suc 1) v) \<le> real (k * (k + 1)) / (d / 2) ^ (k + 2)"
       
  5500                 if "v \<in> ball w (d / 2)" for v
       
  5501       proof -
       
  5502         have "d/2 \<le> cmod (x - v)" using d x that
       
  5503           apply (simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps, clarify)
       
  5504           apply (drule subsetD)
       
  5505            prefer 2 apply blast
       
  5506           apply (metis norm_minus_commute norm_triangle_half_r CollectI)
       
  5507           done
       
  5508         then have "d \<le> cmod (x - v) * 2"
       
  5509           by (simp add: divide_simps)
       
  5510         then have dpow_le: "d ^ (k+2) \<le> (cmod (x - v) * 2) ^ (k+2)"
       
  5511           using \<open>0 < d\<close> order_less_imp_le power_mono by blast
       
  5512         have "x \<noteq> v" using that
       
  5513           using \<open>x \<in> path_image \<gamma>\<close> ball_divide_subset_numeral d by fastforce
       
  5514         then show ?thesis
       
  5515         using \<open>d > 0\<close>
       
  5516         apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc)
       
  5517         using dpow_le
       
  5518         apply (simp add: algebra_simps divide_simps mult_less_0_iff)
       
  5519         done
       
  5520       qed
       
  5521       have ub: "u \<in> ball w (d / 2)"
       
  5522         using uwd by (simp add: dist_commute dist_norm)
       
  5523       have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
       
  5524                   \<le> (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d / 2) ^ k))"
       
  5525         using complex_taylor [OF _ ff1 ff2 _ ub, of w, simplified]
       
  5526         by (simp add: ff_def \<open>0 < d\<close>)
       
  5527       then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
       
  5528                   \<le> (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)"
       
  5529         by (simp add: field_simps)
       
  5530       then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
       
  5531                  / (cmod (u - w) * real k)
       
  5532                   \<le> (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)"
       
  5533         using \<open>k \<noteq> 0\<close> \<open>u \<noteq> w\<close> by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq)
       
  5534       also have "... < e"
       
  5535         using uw_less \<open>0 < d\<close> by (simp add: mult_ac divide_simps)
       
  5536       finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k)))
       
  5537                         / cmod ((u - w) * real k)   <   e"
       
  5538         by (simp add: norm_mult)
       
  5539       have "x \<noteq> u"
       
  5540         using uwd \<open>0 < d\<close> x d by (force simp: dist_norm ball_def norm_minus_commute)
       
  5541       show ?thesis
       
  5542         apply (rule le_less_trans [OF _ e])
       
  5543         using \<open>k \<noteq> 0\<close> \<open>x \<noteq> u\<close>  \<open>u \<noteq> w\<close>
       
  5544         apply (simp add: field_simps norm_divide [symmetric])
       
  5545         done
       
  5546     qed
       
  5547     show ?thesis
       
  5548       unfolding eventually_at
       
  5549       apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI)
       
  5550       apply (force simp: \<open>d > 0\<close> dist_norm that simp del: power_Suc intro: *)
       
  5551       done
       
  5552   qed
       
  5553   have 2: "\<forall>\<^sub>F n in at w.
       
  5554               \<forall>x\<in>path_image \<gamma>.
       
  5555                cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e"
       
  5556           if "0 < e" for e
       
  5557   proof -
       
  5558     have *: "cmod (f' (\<gamma> x) * (inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
       
  5559                         f' (\<gamma> x) / ((\<gamma> x - w) * (\<gamma> x - w) ^ k)) < e"
       
  5560               if ec: "cmod ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
       
  5561                       inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k) < e / C"
       
  5562                  and x: "0 \<le> x" "x \<le> 1"
       
  5563               for u x
       
  5564     proof (cases "(f' (\<gamma> x)) = 0")
       
  5565       case True then show ?thesis by (simp add: \<open>0 < e\<close>)
       
  5566     next
       
  5567       case False
       
  5568       have "cmod (f' (\<gamma> x) * (inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
       
  5569                         f' (\<gamma> x) / ((\<gamma> x - w) * (\<gamma> x - w) ^ k)) =
       
  5570             cmod (f' (\<gamma> x) * ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
       
  5571                              inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k))"
       
  5572         by (simp add: field_simps)
       
  5573       also have "... = cmod (f' (\<gamma> x)) *
       
  5574                        cmod ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
       
  5575                              inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k)"
       
  5576         by (simp add: norm_mult)
       
  5577       also have "... < cmod (f' (\<gamma> x)) * (e/C)"
       
  5578         apply (rule mult_strict_left_mono [OF ec])
       
  5579         using False by simp
       
  5580       also have "... \<le> e" using C
       
  5581         by (metis False \<open>0 < e\<close> frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff)
       
  5582       finally show ?thesis .
       
  5583     qed
       
  5584     show ?thesis
       
  5585       using twom [OF divide_pos_pos [OF that \<open>C > 0\<close>]]   unfolding path_image_def
       
  5586       by (force intro: * elim: eventually_mono)
       
  5587   qed
       
  5588   show "(\<lambda>u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \<gamma>"
       
  5589     by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto
       
  5590   have *: "(\<lambda>n. contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k))
       
  5591            \<midarrow>w\<rightarrow> contour_integral \<gamma> (\<lambda>u. f' u / (u - w) ^ (Suc k))"
       
  5592     by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto
       
  5593   have **: "contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) =
       
  5594               (f u - f w) / (u - w) / k"
       
  5595            if "dist u w < d" for u
       
  5596     apply (rule contour_integral_unique)
       
  5597     apply (simp add: diff_divide_distrib algebra_simps)
       
  5598     apply (rule has_contour_integral_diff; rule has_contour_integral_div; simp add: field_simps; rule int)
       
  5599     apply (metis contra_subsetD d dist_commute mem_ball that)
       
  5600     apply (rule w)
       
  5601     done
       
  5602   show ?thes2
       
  5603     apply (simp add: DERIV_within_iff del: power_Suc)
       
  5604     apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \<open>0 < d\<close> ])
       
  5605     apply (simp add: \<open>k \<noteq> 0\<close> **)
       
  5606     done
       
  5607 qed
       
  5608 
       
  5609 corollary Cauchy_next_derivative_circlepath:
       
  5610   assumes contf: "continuous_on (path_image (circlepath z r)) f"
       
  5611       and int: "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)"
       
  5612       and k: "k \<noteq> 0"
       
  5613       and w: "w \<in> ball z r"
       
  5614     shows "(\<lambda>u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)"
       
  5615            (is "?thes1")
       
  5616       and "(g has_field_derivative (k * contour_integral (circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k)))) (at w)"
       
  5617            (is "?thes2")
       
  5618 proof -
       
  5619   have "r > 0" using w
       
  5620     using ball_eq_empty by fastforce
       
  5621   have wim: "w \<in> ball z r - path_image (circlepath z r)"
       
  5622     using w by (auto simp: dist_norm)
       
  5623   show ?thes1 ?thes2
       
  5624     by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \<bar>r\<bar>"];
       
  5625         auto simp: vector_derivative_circlepath norm_mult)+
       
  5626 qed
       
  5627 
       
  5628 
       
  5629 text\<open> In particular, the first derivative formula.\<close>
       
  5630 
       
  5631 proposition Cauchy_derivative_integral_circlepath:
       
  5632   assumes contf: "continuous_on (cball z r) f"
       
  5633       and holf: "f holomorphic_on ball z r"
       
  5634       and w: "w \<in> ball z r"
       
  5635     shows "(\<lambda>u. f u/(u - w)^2) contour_integrable_on (circlepath z r)"
       
  5636            (is "?thes1")
       
  5637       and "(f has_field_derivative (1 / (2 * of_real pi * \<i>) * contour_integral(circlepath z r) (\<lambda>u. f u / (u - w)^2))) (at w)"
       
  5638            (is "?thes2")
       
  5639 proof -
       
  5640   have [simp]: "r \<ge> 0" using w
       
  5641     using ball_eq_empty by fastforce
       
  5642   have f: "continuous_on (path_image (circlepath z r)) f"
       
  5643     by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def)
       
  5644   have int: "\<And>w. dist z w < r \<Longrightarrow>
       
  5645                  ((\<lambda>u. f u / (u - w)) has_contour_integral (\<lambda>x. 2 * of_real pi * \<i> * f x) w) (circlepath z r)"
       
  5646     by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute)
       
  5647   show ?thes1
       
  5648     apply (simp add: power2_eq_square)
       
  5649     apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified])
       
  5650     apply (blast intro: int)
       
  5651     done
       
  5652   have "((\<lambda>x. 2 * of_real pi * \<i> * f x) has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2)) (at w)"
       
  5653     apply (simp add: power2_eq_square)
       
  5654     apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\<lambda>x. 2 * of_real pi * \<i> * f x", simplified])
       
  5655     apply (blast intro: int)
       
  5656     done
       
  5657   then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2) / (2 * of_real pi * \<i>)) (at w)"
       
  5658     by (rule DERIV_cdivide [where f = "\<lambda>x. 2 * of_real pi * \<i> * f x" and c = "2 * of_real pi * \<i>", simplified])
       
  5659   show ?thes2
       
  5660     by simp (rule fder)
       
  5661 qed
       
  5662 
       
  5663 subsection\<open> Existence of all higher derivatives.\<close>
       
  5664 
       
  5665 proposition derivative_is_holomorphic:
       
  5666   assumes "open s"
       
  5667       and fder: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z)"
       
  5668     shows "f' holomorphic_on s"
       
  5669 proof -
       
  5670   have *: "\<exists>h. (f' has_field_derivative h) (at z)" if "z \<in> s" for z
       
  5671   proof -
       
  5672     obtain r where "r > 0" and r: "cball z r \<subseteq> s"
       
  5673       using open_contains_cball \<open>z \<in> s\<close> \<open>open s\<close> by blast
       
  5674     then have holf_cball: "f holomorphic_on cball z r"
       
  5675       apply (simp add: holomorphic_on_def)
       
  5676       using field_differentiable_at_within field_differentiable_def fder by blast
       
  5677     then have "continuous_on (path_image (circlepath z r)) f"
       
  5678       using \<open>r > 0\<close> by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on])
       
  5679     then have contfpi: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1/(2 * of_real pi*\<i>) * f x)"
       
  5680       by (auto intro: continuous_intros)+
       
  5681     have contf_cball: "continuous_on (cball z r) f" using holf_cball
       
  5682       by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset)
       
  5683     have holf_ball: "f holomorphic_on ball z r" using holf_cball
       
  5684       using ball_subset_cball holomorphic_on_subset by blast
       
  5685     { fix w  assume w: "w \<in> ball z r"
       
  5686       have intf: "(\<lambda>u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r"
       
  5687         by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball])
       
  5688       have fder': "(f has_field_derivative 1 / (2 * of_real pi * \<i>) * contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2))
       
  5689                   (at w)"
       
  5690         by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball])
       
  5691       have f'_eq: "f' w = contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>)"
       
  5692         using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder])
       
  5693       have "((\<lambda>u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \<i>)) has_contour_integral
       
  5694                 contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>))
       
  5695                 (circlepath z r)"
       
  5696         by (rule Cauchy_Integral_Thm.has_contour_integral_div [OF has_contour_integral_integral [OF intf]])
       
  5697       then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral
       
  5698                 contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>))
       
  5699                 (circlepath z r)"
       
  5700         by (simp add: algebra_simps)
       
  5701       then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)"
       
  5702         by (simp add: f'_eq)
       
  5703     } note * = this
       
  5704     show ?thesis
       
  5705       apply (rule exI)
       
  5706       apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified])
       
  5707       apply (simp_all add: \<open>0 < r\<close> * dist_norm)
       
  5708       done
       
  5709   qed
       
  5710   show ?thesis
       
  5711     by (simp add: holomorphic_on_open [OF \<open>open s\<close>] *)
       
  5712 qed
       
  5713 
       
  5714 lemma holomorphic_deriv [holomorphic_intros]:
       
  5715     "\<lbrakk>f holomorphic_on s; open s\<rbrakk> \<Longrightarrow> (deriv f) holomorphic_on s"
       
  5716 by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def)
       
  5717 
       
  5718 lemma analytic_deriv: "f analytic_on s \<Longrightarrow> (deriv f) analytic_on s"
       
  5719   using analytic_on_holomorphic holomorphic_deriv by auto
       
  5720 
       
  5721 lemma holomorphic_higher_deriv [holomorphic_intros]: "\<lbrakk>f holomorphic_on s; open s\<rbrakk> \<Longrightarrow> (deriv ^^ n) f holomorphic_on s"
       
  5722   by (induction n) (auto simp: holomorphic_deriv)
       
  5723 
       
  5724 lemma analytic_higher_deriv: "f analytic_on s \<Longrightarrow> (deriv ^^ n) f analytic_on s"
       
  5725   unfolding analytic_on_def using holomorphic_higher_deriv by blast
       
  5726 
       
  5727 lemma has_field_derivative_higher_deriv:
       
  5728      "\<lbrakk>f holomorphic_on s; open s; x \<in> s\<rbrakk>
       
  5729       \<Longrightarrow> ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)"
       
  5730 by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply
       
  5731          funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def)
       
  5732 
       
  5733 lemma valid_path_compose_holomorphic:
       
  5734   assumes "valid_path g" and holo:"f holomorphic_on s" and "open s" "path_image g \<subseteq> s"
       
  5735   shows "valid_path (f o g)"
       
  5736 proof (rule valid_path_compose[OF \<open>valid_path g\<close>])
       
  5737   fix x assume "x \<in> path_image g"
       
  5738   then show "\<exists>f'. (f has_field_derivative f') (at x)"
       
  5739     using holo holomorphic_on_open[OF \<open>open s\<close>] \<open>path_image g \<subseteq> s\<close> by auto
       
  5740 next
       
  5741   have "deriv f holomorphic_on s"
       
  5742     using holomorphic_deriv holo \<open>open s\<close> by auto
       
  5743   then show "continuous_on (path_image g) (deriv f)"
       
  5744     using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto
       
  5745 qed
       
  5746 
       
  5747 
       
  5748 subsection\<open> Morera's theorem.\<close>
       
  5749 
       
  5750 lemma Morera_local_triangle_ball:
       
  5751   assumes "\<And>z. z \<in> s
       
  5752           \<Longrightarrow> \<exists>e a. 0 < e \<and> z \<in> ball a e \<and> continuous_on (ball a e) f \<and>
       
  5753                     (\<forall>b c. closed_segment b c \<subseteq> ball a e
       
  5754                            \<longrightarrow> contour_integral (linepath a b) f +
       
  5755                                contour_integral (linepath b c) f +
       
  5756                                contour_integral (linepath c a) f = 0)"
       
  5757   shows "f analytic_on s"
       
  5758 proof -
       
  5759   { fix z  assume "z \<in> s"
       
  5760     with assms obtain e a where
       
  5761             "0 < e" and z: "z \<in> ball a e" and contf: "continuous_on (ball a e) f"
       
  5762         and 0: "\<And>b c. closed_segment b c \<subseteq> ball a e
       
  5763                       \<Longrightarrow> contour_integral (linepath a b) f +
       
  5764                           contour_integral (linepath b c) f +
       
  5765                           contour_integral (linepath c a) f = 0"
       
  5766       by fastforce
       
  5767     have az: "dist a z < e" using mem_ball z by blast
       
  5768     have sb_ball: "ball z (e - dist a z) \<subseteq> ball a e"
       
  5769       by (simp add: dist_commute ball_subset_ball_iff)
       
  5770     have "\<exists>e>0. f holomorphic_on ball z e"
       
  5771       apply (rule_tac x="e - dist a z" in exI)
       
  5772       apply (simp add: az)
       
  5773       apply (rule holomorphic_on_subset [OF _ sb_ball])
       
  5774       apply (rule derivative_is_holomorphic[OF open_ball])
       
  5775       apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a])
       
  5776          apply (simp_all add: 0 \<open>0 < e\<close>)
       
  5777       apply (meson \<open>0 < e\<close> centre_in_ball convex_ball convex_contains_segment mem_ball)
       
  5778       done
       
  5779   }
       
  5780   then show ?thesis
       
  5781     by (simp add: analytic_on_def)
       
  5782 qed
       
  5783 
       
  5784 lemma Morera_local_triangle:
       
  5785   assumes "\<And>z. z \<in> s
       
  5786           \<Longrightarrow> \<exists>t. open t \<and> z \<in> t \<and> continuous_on t f \<and>
       
  5787                   (\<forall>a b c. convex hull {a,b,c} \<subseteq> t
       
  5788                               \<longrightarrow> contour_integral (linepath a b) f +
       
  5789                                   contour_integral (linepath b c) f +
       
  5790                                   contour_integral (linepath c a) f = 0)"
       
  5791   shows "f analytic_on s"
       
  5792 proof -
       
  5793   { fix z  assume "z \<in> s"
       
  5794     with assms obtain t where
       
  5795             "open t" and z: "z \<in> t" and contf: "continuous_on t f"
       
  5796         and 0: "\<And>a b c. convex hull {a,b,c} \<subseteq> t
       
  5797                       \<Longrightarrow> contour_integral (linepath a b) f +
       
  5798                           contour_integral (linepath b c) f +
       
  5799                           contour_integral (linepath c a) f = 0"
       
  5800       by force
       
  5801     then obtain e where "e>0" and e: "ball z e \<subseteq> t"
       
  5802       using open_contains_ball by blast
       
  5803     have [simp]: "continuous_on (ball z e) f" using contf
       
  5804       using continuous_on_subset e by blast
       
  5805     have "\<exists>e a. 0 < e \<and>
       
  5806                z \<in> ball a e \<and>
       
  5807                continuous_on (ball a e) f \<and>
       
  5808                (\<forall>b c. closed_segment b c \<subseteq> ball a e \<longrightarrow>
       
  5809                       contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)"
       
  5810       apply (rule_tac x=e in exI)
       
  5811       apply (rule_tac x=z in exI)
       
  5812       apply (simp add: \<open>e > 0\<close>, clarify)
       
  5813       apply (rule 0)
       
  5814       apply (meson z \<open>0 < e\<close> centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset)
       
  5815       done
       
  5816   }
       
  5817   then show ?thesis
       
  5818     by (simp add: Morera_local_triangle_ball)
       
  5819 qed
       
  5820 
       
  5821 proposition Morera_triangle:
       
  5822     "\<lbrakk>continuous_on s f; open s;
       
  5823       \<And>a b c. convex hull {a,b,c} \<subseteq> s
       
  5824               \<longrightarrow> contour_integral (linepath a b) f +
       
  5825                   contour_integral (linepath b c) f +
       
  5826                   contour_integral (linepath c a) f = 0\<rbrakk>
       
  5827      \<Longrightarrow> f analytic_on s"
       
  5828   using Morera_local_triangle by blast
       
  5829 
       
  5830 
       
  5831 
       
  5832 subsection\<open> Combining theorems for higher derivatives including Leibniz rule.\<close>
       
  5833 
       
  5834 lemma higher_deriv_linear [simp]:
       
  5835     "(deriv ^^ n) (\<lambda>w. c*w) = (\<lambda>z. if n = 0 then c*z else if n = 1 then c else 0)"
       
  5836   by (induction n) (auto simp: deriv_const deriv_linear)
       
  5837 
       
  5838 lemma higher_deriv_const [simp]: "(deriv ^^ n) (\<lambda>w. c) = (\<lambda>w. if n=0 then c else 0)"
       
  5839   by (induction n) (auto simp: deriv_const)
       
  5840 
       
  5841 lemma higher_deriv_ident [simp]:
       
  5842      "(deriv ^^ n) (\<lambda>w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)"
       
  5843   apply (induction n, simp)
       
  5844   apply (metis higher_deriv_linear lambda_one)
       
  5845   done
       
  5846 
       
  5847 corollary higher_deriv_id [simp]:
       
  5848      "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)"
       
  5849   by (simp add: id_def)
       
  5850 
       
  5851 lemma has_complex_derivative_funpow_1:
       
  5852      "\<lbrakk>(f has_field_derivative 1) (at z); f z = z\<rbrakk> \<Longrightarrow> (f^^n has_field_derivative 1) (at z)"
       
  5853   apply (induction n)
       
  5854   apply auto
       
  5855   apply (metis DERIV_ident DERIV_transform_at id_apply zero_less_one)
       
  5856   by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral)
       
  5857 
       
  5858 proposition higher_deriv_uminus:
       
  5859   assumes "f holomorphic_on s" "open s" and z: "z \<in> s"
       
  5860     shows "(deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)"
       
  5861 using z
       
  5862 proof (induction n arbitrary: z)
       
  5863   case 0 then show ?case by simp
       
  5864 next
       
  5865   case (Suc n z)
       
  5866   have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
       
  5867     using Suc.prems assms has_field_derivative_higher_deriv by auto
       
  5868   show ?case
       
  5869     apply simp
       
  5870     apply (rule DERIV_imp_deriv)
       
  5871     apply (rule DERIV_transform_within_open [of "\<lambda>w. -((deriv ^^ n) f w)"])
       
  5872     apply (rule derivative_eq_intros | rule * refl assms Suc)+
       
  5873     apply (simp add: Suc)
       
  5874     done
       
  5875 qed
       
  5876 
       
  5877 proposition higher_deriv_add:
       
  5878   fixes z::complex
       
  5879   assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
       
  5880     shows "(deriv ^^ n) (\<lambda>w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z"
       
  5881 using z
       
  5882 proof (induction n arbitrary: z)
       
  5883   case 0 then show ?case by simp
       
  5884 next
       
  5885   case (Suc n z)
       
  5886   have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
       
  5887           "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)"
       
  5888     using Suc.prems assms has_field_derivative_higher_deriv by auto
       
  5889   show ?case
       
  5890     apply simp
       
  5891     apply (rule DERIV_imp_deriv)
       
  5892     apply (rule DERIV_transform_within_open [of "\<lambda>w. (deriv ^^ n) f w + (deriv ^^ n) g w"])
       
  5893     apply (rule derivative_eq_intros | rule * refl assms Suc)+
       
  5894     apply (simp add: Suc)
       
  5895     done
       
  5896 qed
       
  5897 
       
  5898 corollary higher_deriv_diff:
       
  5899   fixes z::complex
       
  5900   assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
       
  5901     shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z"
       
  5902   apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add)
       
  5903   apply (subst higher_deriv_add)
       
  5904   using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus)
       
  5905   done
       
  5906 
       
  5907 
       
  5908 lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))"
       
  5909   by (cases k) simp_all
       
  5910 
       
  5911 proposition higher_deriv_mult:
       
  5912   fixes z::complex
       
  5913   assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
       
  5914     shows "(deriv ^^ n) (\<lambda>w. f w * g w) z =
       
  5915            (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)"
       
  5916 using z
       
  5917 proof (induction n arbitrary: z)
       
  5918   case 0 then show ?case by simp
       
  5919 next
       
  5920   case (Suc n z)
       
  5921   have *: "\<And>n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
       
  5922           "\<And>n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)"
       
  5923     using Suc.prems assms has_field_derivative_higher_deriv by auto
       
  5924   have sumeq: "(\<Sum>i = 0..n.
       
  5925                of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) =
       
  5926             g z * deriv ((deriv ^^ n) f) z + (\<Sum>i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))"
       
  5927     apply (simp add: bb algebra_simps setsum.distrib)
       
  5928     apply (subst (4) setsum_Suc_reindex)
       
  5929     apply (auto simp: algebra_simps Suc_diff_le intro: setsum.cong)
       
  5930     done
       
  5931   show ?case
       
  5932     apply (simp only: funpow.simps o_apply)
       
  5933     apply (rule DERIV_imp_deriv)
       
  5934     apply (rule DERIV_transform_within_open
       
  5935              [of "\<lambda>w. (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"])
       
  5936     apply (simp add: algebra_simps)
       
  5937     apply (rule DERIV_cong [OF DERIV_setsum])
       
  5938     apply (rule DERIV_cmult)
       
  5939     apply (auto simp: intro: DERIV_mult * sumeq \<open>open s\<close> Suc.prems Suc.IH [symmetric])
       
  5940     done
       
  5941 qed
       
  5942 
       
  5943 
       
  5944 proposition higher_deriv_transform_within_open:
       
  5945   fixes z::complex
       
  5946   assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
       
  5947       and fg: "\<And>w. w \<in> s \<Longrightarrow> f w = g w"
       
  5948     shows "(deriv ^^ i) f z = (deriv ^^ i) g z"
       
  5949 using z
       
  5950 by (induction i arbitrary: z)
       
  5951    (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms)
       
  5952 
       
  5953 proposition higher_deriv_compose_linear:
       
  5954   fixes z::complex
       
  5955   assumes f: "f holomorphic_on t" and s: "open s" and t: "open t" and z: "z \<in> s"
       
  5956       and fg: "\<And>w. w \<in> s \<Longrightarrow> u * w \<in> t"
       
  5957     shows "(deriv ^^ n) (\<lambda>w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)"
       
  5958 using z
       
  5959 proof (induction n arbitrary: z)
       
  5960   case 0 then show ?case by simp
       
  5961 next
       
  5962   case (Suc n z)
       
  5963   have holo0: "f holomorphic_on op * u ` s"
       
  5964     by (meson fg f holomorphic_on_subset image_subset_iff)
       
  5965   have holo1: "(\<lambda>w. f (u * w)) holomorphic_on s"
       
  5966     apply (rule holomorphic_on_compose [where g=f, unfolded o_def])
       
  5967     apply (rule holo0 holomorphic_intros)+
       
  5968     done
       
  5969   have holo2: "(\<lambda>z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on s"
       
  5970     apply (rule holomorphic_intros)+
       
  5971     apply (rule holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def])
       
  5972     apply (rule holomorphic_intros)
       
  5973     apply (rule holomorphic_on_subset [where s=t])
       
  5974     apply (rule holomorphic_intros assms)+
       
  5975     apply (blast intro: fg)
       
  5976     done
       
  5977   have "deriv ((deriv ^^ n) (\<lambda>w. f (u * w))) z = deriv (\<lambda>z. u^n * (deriv ^^ n) f (u*z)) z"
       
  5978     apply (rule complex_derivative_transform_within_open [OF _ holo2 s Suc.prems])
       
  5979     apply (rule holomorphic_higher_deriv [OF holo1 s])
       
  5980     apply (simp add: Suc.IH)
       
  5981     done
       
  5982   also have "... = u^n * deriv (\<lambda>z. (deriv ^^ n) f (u * z)) z"
       
  5983     apply (rule deriv_cmult)
       
  5984     apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems])
       
  5985     apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and t=t, unfolded o_def])
       
  5986     apply (simp add: analytic_on_linear)
       
  5987     apply (simp add: analytic_on_open f holomorphic_higher_deriv t)
       
  5988     apply (blast intro: fg)
       
  5989     done
       
  5990   also have "... = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)"
       
  5991       apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "op*u", unfolded o_def])
       
  5992       apply (rule derivative_intros)
       
  5993       using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv t apply blast
       
  5994       apply (simp add: deriv_linear)
       
  5995       done
       
  5996   finally show ?case
       
  5997     by simp
       
  5998 qed
       
  5999 
       
  6000 lemma higher_deriv_add_at:
       
  6001   assumes "f analytic_on {z}" "g analytic_on {z}"
       
  6002     shows "(deriv ^^ n) (\<lambda>w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z"
       
  6003 proof -
       
  6004   have "f analytic_on {z} \<and> g analytic_on {z}"
       
  6005     using assms by blast
       
  6006   with higher_deriv_add show ?thesis
       
  6007     by (auto simp: analytic_at_two)
       
  6008 qed
       
  6009 
       
  6010 lemma higher_deriv_diff_at:
       
  6011   assumes "f analytic_on {z}" "g analytic_on {z}"
       
  6012     shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z"
       
  6013 proof -
       
  6014   have "f analytic_on {z} \<and> g analytic_on {z}"
       
  6015     using assms by blast
       
  6016   with higher_deriv_diff show ?thesis
       
  6017     by (auto simp: analytic_at_two)
       
  6018 qed
       
  6019 
       
  6020 lemma higher_deriv_uminus_at:
       
  6021    "f analytic_on {z}  \<Longrightarrow> (deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)"
       
  6022   using higher_deriv_uminus
       
  6023     by (auto simp: analytic_at)
       
  6024 
       
  6025 lemma higher_deriv_mult_at:
       
  6026   assumes "f analytic_on {z}" "g analytic_on {z}"
       
  6027     shows "(deriv ^^ n) (\<lambda>w. f w * g w) z =
       
  6028            (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)"
       
  6029 proof -
       
  6030   have "f analytic_on {z} \<and> g analytic_on {z}"
       
  6031     using assms by blast
       
  6032   with higher_deriv_mult show ?thesis
       
  6033     by (auto simp: analytic_at_two)
       
  6034 qed
       
  6035 
       
  6036 
       
  6037 text\<open> Nonexistence of isolated singularities and a stronger integral formula.\<close>
       
  6038 
       
  6039 proposition no_isolated_singularity:
       
  6040   fixes z::complex
       
  6041   assumes f: "continuous_on s f" and holf: "f holomorphic_on (s - k)" and s: "open s" and k: "finite k"
       
  6042     shows "f holomorphic_on s"
       
  6043 proof -
       
  6044   { fix z
       
  6045     assume "z \<in> s" and cdf: "\<And>x. x\<in>s - k \<Longrightarrow> f field_differentiable at x"
       
  6046     have "f field_differentiable at z"
       
  6047     proof (cases "z \<in> k")
       
  6048       case False then show ?thesis by (blast intro: cdf \<open>z \<in> s\<close>)
       
  6049     next
       
  6050       case True
       
  6051       with finite_set_avoid [OF k, of z]
       
  6052       obtain d where "d>0" and d: "\<And>x. \<lbrakk>x\<in>k; x \<noteq> z\<rbrakk> \<Longrightarrow> d \<le> dist z x"
       
  6053         by blast
       
  6054       obtain e where "e>0" and e: "ball z e \<subseteq> s"
       
  6055         using  s \<open>z \<in> s\<close> by (force simp add: open_contains_ball)
       
  6056       have fde: "continuous_on (ball z (min d e)) f"
       
  6057         by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI)
       
  6058       have "\<exists>g. \<forall>w \<in> ball z (min d e). (g has_field_derivative f w) (at w within ball z (min d e))"
       
  6059         apply (rule contour_integral_convex_primitive [OF convex_ball fde])
       
  6060         apply (rule Cauchy_theorem_triangle_cofinite [OF _ k])
       
  6061          apply (metis continuous_on_subset [OF fde] closed_segment_subset convex_ball starlike_convex_subset)
       
  6062         apply (rule cdf)
       
  6063         apply (metis Diff_iff Int_iff ball_min_Int bot_least contra_subsetD convex_ball e insert_subset
       
  6064                interior_mono interior_subset subset_hull)
       
  6065         done
       
  6066       then have "f holomorphic_on ball z (min d e)"
       
  6067         by (metis open_ball at_within_open derivative_is_holomorphic)
       
  6068       then show ?thesis
       
  6069         unfolding holomorphic_on_def
       
  6070         by (metis open_ball \<open>0 < d\<close> \<open>0 < e\<close> at_within_open centre_in_ball min_less_iff_conj)
       
  6071     qed
       
  6072   }
       
  6073   with holf s k show ?thesis
       
  6074     by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric])
       
  6075 qed
       
  6076 
       
  6077 proposition Cauchy_integral_formula_convex:
       
  6078     assumes s: "convex s" and k: "finite k" and contf: "continuous_on s f"
       
  6079         and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)"
       
  6080         and z: "z \<in> interior s" and vpg: "valid_path \<gamma>"
       
  6081         and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  6082       shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
       
  6083   apply (rule Cauchy_integral_formula_weak [OF s finite.emptyI contf])
       
  6084   apply (simp add: holomorphic_on_open [symmetric] field_differentiable_def)
       
  6085   using no_isolated_singularity [where s = "interior s"]
       
  6086   apply (metis k contf fcd holomorphic_on_open field_differentiable_def continuous_on_subset
       
  6087                has_field_derivative_at_within holomorphic_on_def interior_subset open_interior)
       
  6088   using assms
       
  6089   apply auto
       
  6090   done
       
  6091 
       
  6092 
       
  6093 text\<open> Formula for higher derivatives.\<close>
       
  6094 
       
  6095 proposition Cauchy_has_contour_integral_higher_derivative_circlepath:
       
  6096   assumes contf: "continuous_on (cball z r) f"
       
  6097       and holf: "f holomorphic_on ball z r"
       
  6098       and w: "w \<in> ball z r"
       
  6099     shows "((\<lambda>u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \<i>) / (fact k) * (deriv ^^ k) f w))
       
  6100            (circlepath z r)"
       
  6101 using w
       
  6102 proof (induction k arbitrary: w)
       
  6103   case 0 then show ?case
       
  6104     using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm)
       
  6105 next
       
  6106   case (Suc k)
       
  6107   have [simp]: "r > 0" using w
       
  6108     using ball_eq_empty by fastforce
       
  6109   have f: "continuous_on (path_image (circlepath z r)) f"
       
  6110     by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def less_imp_le)
       
  6111   obtain X where X: "((\<lambda>u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)"
       
  6112     using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems]
       
  6113     by (auto simp: contour_integrable_on_def)
       
  6114   then have con: "contour_integral (circlepath z r) ((\<lambda>u. f u / (u - w) ^ Suc (Suc k))) = X"
       
  6115     by (rule contour_integral_unique)
       
  6116   have "\<And>n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)"
       
  6117     using Suc.prems assms has_field_derivative_higher_deriv by auto
       
  6118   then have dnf_diff: "\<And>n. (deriv ^^ n) f field_differentiable (at w)"
       
  6119     by (force simp add: field_differentiable_def)
       
  6120   have "deriv (\<lambda>w. complex_of_real (2 * pi) * \<i> / (fact k) * (deriv ^^ k) f w) w =
       
  6121           of_nat (Suc k) * contour_integral (circlepath z r) (\<lambda>u. f u / (u - w) ^ Suc (Suc k))"
       
  6122     by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems])
       
  6123   also have "... = of_nat (Suc k) * X"
       
  6124     by (simp only: con)
       
  6125   finally have "deriv (\<lambda>w. ((2 * pi) * \<i> / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" .
       
  6126   then have "((2 * pi) * \<i> / (fact k)) * deriv (\<lambda>w. (deriv ^^ k) f w) w = of_nat (Suc k) * X"
       
  6127     by (metis deriv_cmult dnf_diff)
       
  6128   then have "deriv (\<lambda>w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \<i> / (fact k))"
       
  6129     by (simp add: field_simps)
       
  6130   then show ?case
       
  6131   using of_nat_eq_0_iff X by fastforce
       
  6132 qed
       
  6133 
       
  6134 proposition Cauchy_higher_derivative_integral_circlepath:
       
  6135   assumes contf: "continuous_on (cball z r) f"
       
  6136       and holf: "f holomorphic_on ball z r"
       
  6137       and w: "w \<in> ball z r"
       
  6138     shows "(\<lambda>u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)"
       
  6139            (is "?thes1")
       
  6140       and "(deriv ^^ k) f w = (fact k) / (2 * pi * \<i>) * contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k))"
       
  6141            (is "?thes2")
       
  6142 proof -
       
  6143   have *: "((\<lambda>u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \<i> / (fact k) * (deriv ^^ k) f w)
       
  6144            (circlepath z r)"
       
  6145     using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms]
       
  6146     by simp
       
  6147   show ?thes1 using *
       
  6148     using contour_integrable_on_def by blast
       
  6149   show ?thes2
       
  6150     unfolding contour_integral_unique [OF *] by (simp add: divide_simps)
       
  6151 qed
       
  6152 
       
  6153 corollary Cauchy_contour_integral_circlepath:
       
  6154   assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \<in> ball z r"
       
  6155     shows "contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k)) = (2 * pi * \<i>) * (deriv ^^ k) f w / (fact k)"
       
  6156 by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms])
       
  6157 
       
  6158 corollary Cauchy_contour_integral_circlepath_2:
       
  6159   assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \<in> ball z r"
       
  6160     shows "contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^2) = (2 * pi * \<i>) * deriv f w"
       
  6161   using Cauchy_contour_integral_circlepath [OF assms, of 1]
       
  6162   by (simp add: power2_eq_square)
       
  6163 
       
  6164 
       
  6165 subsection\<open>A holomorphic function is analytic, i.e. has local power series.\<close>
       
  6166 
       
  6167 theorem holomorphic_power_series:
       
  6168   assumes holf: "f holomorphic_on ball z r"
       
  6169       and w: "w \<in> ball z r"
       
  6170     shows "((\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
       
  6171 proof -
       
  6172   have fh': "f holomorphic_on cball z ((r + dist w z) / 2)"
       
  6173      apply (rule holomorphic_on_subset [OF holf])
       
  6174      apply (clarsimp simp del: divide_const_simps)
       
  6175      apply (metis add.commute dist_commute le_less_trans mem_ball real_gt_half_sum w)
       
  6176      done
       
  6177   \<comment>\<open>Replacing @{term r} and the original (weak) premises\<close>
       
  6178   obtain r where "0 < r" and holfc: "f holomorphic_on cball z r" and w: "w \<in> ball z r"
       
  6179     apply (rule that [of "(r + dist w z) / 2"])
       
  6180       apply (simp_all add: fh')
       
  6181      apply (metis add_0_iff ball_eq_empty dist_nz dist_self empty_iff not_less pos_add_strict w)
       
  6182     apply (metis add_less_cancel_right dist_commute mem_ball mult_2_right w)
       
  6183     done
       
  6184   then have holf: "f holomorphic_on ball z r" and contf: "continuous_on (cball z r) f"
       
  6185     using ball_subset_cball holomorphic_on_subset apply blast
       
  6186     by (simp add: holfc holomorphic_on_imp_continuous_on)
       
  6187   have cint: "\<And>k. (\<lambda>u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r"
       
  6188     apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf])
       
  6189     apply (simp add: \<open>0 < r\<close>)
       
  6190     done
       
  6191   obtain B where "0 < B" and B: "\<And>u. u \<in> cball z r \<Longrightarrow> norm(f u) \<le> B"
       
  6192     by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI)
       
  6193   obtain k where k: "0 < k" "k \<le> r" and wz_eq: "norm(w - z) = r - k"
       
  6194              and kle: "\<And>u. norm(u - z) = r \<Longrightarrow> k \<le> norm(u - w)"
       
  6195     apply (rule_tac k = "r - dist z w" in that)
       
  6196     using w
       
  6197     apply (auto simp: dist_norm norm_minus_commute)
       
  6198     by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
       
  6199   have *: "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>path_image (circlepath z r).
       
  6200                 norm ((\<Sum>k<n. (w - z) ^ k * (f x / (x - z) ^ Suc k)) - f x / (x - w)) < e"
       
  6201           if "0 < e" for e
       
  6202   proof -
       
  6203     have rr: "0 \<le> (r - k) / r" "(r - k) / r < 1" using  k by auto
       
  6204     obtain n where n: "((r - k) / r) ^ n < e / B * k"
       
  6205       using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \<open>0 < e\<close> \<open>0 < B\<close> k by force
       
  6206     have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) - f u / (u - w)) < e"
       
  6207          if "n \<le> N" and r: "r = dist z u"  for N u
       
  6208     proof -
       
  6209       have N: "((r - k) / r) ^ N < e / B * k"
       
  6210         apply (rule le_less_trans [OF power_decreasing n])
       
  6211         using  \<open>n \<le> N\<close> k by auto
       
  6212       have u [simp]: "(u \<noteq> z) \<and> (u \<noteq> w)"
       
  6213         using \<open>0 < r\<close> r w by auto
       
  6214       have wzu_not1: "(w - z) / (u - z) \<noteq> 1"
       
  6215         by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w)
       
  6216       have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) * (u - w) - f u)
       
  6217             = norm ((\<Sum>k<N. (((w - z) / (u - z)) ^ k)) * f u * (u - w) / (u - z) - f u)"
       
  6218         unfolding setsum_left_distrib setsum_divide_distrib power_divide by (simp add: algebra_simps)
       
  6219       also have "... = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)"
       
  6220         using \<open>0 < B\<close>
       
  6221         apply (auto simp: geometric_sum [OF wzu_not1])
       
  6222         apply (simp add: field_simps norm_mult [symmetric])
       
  6223         done
       
  6224       also have "... = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)"
       
  6225         using \<open>0 < r\<close> r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute)
       
  6226       also have "... = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)"
       
  6227         by (simp add: algebra_simps)
       
  6228       also have "... = norm (w - z) ^ N * norm (f u) / r ^ N"
       
  6229         by (simp add: norm_mult norm_power norm_minus_commute)
       
  6230       also have "... \<le> (((r - k)/r)^N) * B"
       
  6231         using \<open>0 < r\<close> w k
       
  6232         apply (simp add: divide_simps)
       
  6233         apply (rule mult_mono [OF power_mono])
       
  6234         apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r)
       
  6235         done
       
  6236       also have "... < e * k"
       
  6237         using \<open>0 < B\<close> N by (simp add: divide_simps)
       
  6238       also have "... \<le> e * norm (u - w)"
       
  6239         using r kle \<open>0 < e\<close> by (simp add: dist_commute dist_norm)
       
  6240       finally show ?thesis
       
  6241         by (simp add: divide_simps norm_divide del: power_Suc)
       
  6242     qed
       
  6243     with \<open>0 < r\<close> show ?thesis
       
  6244       by (auto simp: mult_ac less_imp_le eventually_sequentially Ball_def)
       
  6245   qed
       
  6246   have eq: "\<forall>\<^sub>F x in sequentially.
       
  6247              contour_integral (circlepath z r) (\<lambda>u. \<Sum>k<x. (w - z) ^ k * (f u / (u - z) ^ Suc k)) =
       
  6248              (\<Sum>k<x. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z) ^ k)"
       
  6249     apply (rule eventuallyI)
       
  6250     apply (subst contour_integral_setsum, simp)
       
  6251     using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps)
       
  6252     apply (simp only: contour_integral_lmul cint algebra_simps)
       
  6253     done
       
  6254   have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
       
  6255         sums contour_integral (circlepath z r) (\<lambda>u. f u/(u - w))"
       
  6256     unfolding sums_def
       
  6257     apply (rule Lim_transform_eventually [OF eq])
       
  6258     apply (rule contour_integral_uniform_limit_circlepath [OF eventuallyI *])
       
  6259     apply (rule contour_integrable_setsum, simp)
       
  6260     apply (rule contour_integrable_lmul)
       
  6261     apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf])
       
  6262     using \<open>0 < r\<close>
       
  6263     apply auto
       
  6264     done
       
  6265   then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
       
  6266              sums (2 * of_real pi * \<i> * f w)"
       
  6267     using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]])
       
  6268   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)))
       
  6269             sums ((2 * of_real pi * \<i> * f w) / (\<i> * (complex_of_real pi * 2)))"
       
  6270     by (rule sums_divide)
       
  6271   then have "(\<lambda>n. (w - z) ^ n * contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc n) / (\<i> * (of_real pi * 2)))
       
  6272             sums f w"
       
  6273     by (simp add: field_simps)
       
  6274   then show ?thesis
       
  6275     by (simp add: field_simps \<open>0 < r\<close> Cauchy_higher_derivative_integral_circlepath [OF contf holf])
       
  6276 qed
       
  6277 
       
  6278 
       
  6279 subsection\<open>The Liouville theorem and the Fundamental Theorem of Algebra.\<close>
       
  6280 
       
  6281 text\<open> These weak Liouville versions don't even need the derivative formula.\<close>
       
  6282 
       
  6283 lemma Liouville_weak_0:
       
  6284   assumes holf: "f holomorphic_on UNIV" and inf: "(f \<longlongrightarrow> 0) at_infinity"
       
  6285     shows "f z = 0"
       
  6286 proof (rule ccontr)
       
  6287   assume fz: "f z \<noteq> 0"
       
  6288   with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"]
       
  6289   obtain B where B: "\<And>x. B \<le> cmod x \<Longrightarrow> norm (f x) * 2 < cmod (f z)"
       
  6290     by (auto simp: dist_norm)
       
  6291   define R where "R = 1 + \<bar>B\<bar> + norm z"
       
  6292   have "R > 0" unfolding R_def
       
  6293   proof -
       
  6294     have "0 \<le> cmod z + \<bar>B\<bar>"
       
  6295       by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def)
       
  6296     then show "0 < 1 + \<bar>B\<bar> + cmod z"
       
  6297       by linarith
       
  6298   qed
       
  6299   have *: "((\<lambda>u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \<i> * f z) (circlepath z R)"
       
  6300     apply (rule Cauchy_integral_circlepath)
       
  6301     using \<open>R > 0\<close> apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+
       
  6302     done
       
  6303   have "cmod (x - z) = R \<Longrightarrow> cmod (f x) * 2 \<le> cmod (f z)" for x
       
  6304     apply (simp add: R_def)
       
  6305     apply (rule less_imp_le)
       
  6306     apply (rule B)
       
  6307     using norm_triangle_ineq4 [of x z]
       
  6308     apply (auto simp:)
       
  6309     done
       
  6310   with  \<open>R > 0\<close> fz show False
       
  6311     using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"]
       
  6312     by (auto simp: norm_mult norm_divide divide_simps)
       
  6313 qed
       
  6314 
       
  6315 proposition Liouville_weak:
       
  6316   assumes "f holomorphic_on UNIV" and "(f \<longlongrightarrow> l) at_infinity"
       
  6317     shows "f z = l"
       
  6318   using Liouville_weak_0 [of "\<lambda>z. f z - l"]
       
  6319   by (simp add: assms holomorphic_on_const holomorphic_on_diff LIM_zero)
       
  6320 
       
  6321 
       
  6322 proposition Liouville_weak_inverse:
       
  6323   assumes "f holomorphic_on UNIV" and unbounded: "\<And>B. eventually (\<lambda>x. norm (f x) \<ge> B) at_infinity"
       
  6324     obtains z where "f z = 0"
       
  6325 proof -
       
  6326   { assume f: "\<And>z. f z \<noteq> 0"
       
  6327     have 1: "(\<lambda>x. 1 / f x) holomorphic_on UNIV"
       
  6328       by (simp add: holomorphic_on_divide holomorphic_on_const assms f)
       
  6329     have 2: "((\<lambda>x. 1 / f x) \<longlongrightarrow> 0) at_infinity"
       
  6330       apply (rule tendstoI [OF eventually_mono])
       
  6331       apply (rule_tac B="2/e" in unbounded)
       
  6332       apply (simp add: dist_norm norm_divide divide_simps mult_ac)
       
  6333       done
       
  6334     have False
       
  6335       using Liouville_weak_0 [OF 1 2] f by simp
       
  6336   }
       
  6337   then show ?thesis
       
  6338     using that by blast
       
  6339 qed
       
  6340 
       
  6341 
       
  6342 text\<open> In particular we get the Fundamental Theorem of Algebra.\<close>
       
  6343 
       
  6344 theorem fundamental_theorem_of_algebra:
       
  6345     fixes a :: "nat \<Rightarrow> complex"
       
  6346   assumes "a 0 = 0 \<or> (\<exists>i \<in> {1..n}. a i \<noteq> 0)"
       
  6347   obtains z where "(\<Sum>i\<le>n. a i * z^i) = 0"
       
  6348 using assms
       
  6349 proof (elim disjE bexE)
       
  6350   assume "a 0 = 0" then show ?thesis
       
  6351     by (auto simp: that [of 0])
       
  6352 next
       
  6353   fix i
       
  6354   assume i: "i \<in> {1..n}" and nz: "a i \<noteq> 0"
       
  6355   have 1: "(\<lambda>z. \<Sum>i\<le>n. a i * z^i) holomorphic_on UNIV"
       
  6356     by (rule holomorphic_intros)+
       
  6357   show ?thesis
       
  6358     apply (rule Liouville_weak_inverse [OF 1])
       
  6359     apply (rule polyfun_extremal)
       
  6360     apply (rule nz)
       
  6361     using i that
       
  6362     apply (auto simp:)
       
  6363     done
       
  6364 qed
       
  6365 
       
  6366 
       
  6367 subsection\<open> Weierstrass convergence theorem.\<close>
       
  6368 
       
  6369 proposition holomorphic_uniform_limit:
       
  6370   assumes cont: "eventually (\<lambda>n. continuous_on (cball z r) (f n) \<and> (f n) holomorphic_on ball z r) F"
       
  6371       and lim: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> cball z r. norm(f n x - g x) < e) F"
       
  6372       and F:  "~ trivial_limit F"
       
  6373   obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r"
       
  6374 proof (cases r "0::real" rule: linorder_cases)
       
  6375   case less then show ?thesis by (force simp add: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that)
       
  6376 next
       
  6377   case equal then show ?thesis
       
  6378     by (force simp add: holomorphic_on_def continuous_on_sing intro: that)
       
  6379 next
       
  6380   case greater
       
  6381   have contg: "continuous_on (cball z r) g"
       
  6382     using cont
       
  6383     by (fastforce simp: eventually_conj_iff dist_norm intro: eventually_mono [OF lim] continuous_uniform_limit [OF F])
       
  6384   have 1: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1 / (2 * complex_of_real pi * \<i>) * g x)"
       
  6385     apply (rule continuous_intros continuous_on_subset [OF contg])+
       
  6386     using \<open>0 < r\<close> by auto
       
  6387   have 2: "((\<lambda>u. 1 / (2 * of_real pi * \<i>) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)"
       
  6388        if w: "w \<in> ball z r" for w
       
  6389   proof -
       
  6390     define d where "d = (r - norm(w - z))"
       
  6391     have "0 < d"  "d \<le> r" using w by (auto simp: norm_minus_commute d_def dist_norm)
       
  6392     have dle: "\<And>u. cmod (z - u) = r \<Longrightarrow> d \<le> cmod (u - w)"
       
  6393       unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
       
  6394     have ev_int: "\<forall>\<^sub>F n in F. (\<lambda>u. f n u / (u - w)) contour_integrable_on circlepath z r"
       
  6395       apply (rule eventually_mono [OF cont])
       
  6396       using w
       
  6397       apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified])
       
  6398       done
       
  6399     have ev_less: "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image (circlepath z r). cmod (f n x / (x - w) - g x / (x - w)) < e"
       
  6400          if "e > 0" for e
       
  6401       using greater \<open>0 < d\<close> \<open>0 < e\<close>
       
  6402       apply (simp add: norm_divide diff_divide_distrib [symmetric] divide_simps)
       
  6403       apply (rule_tac e1="e * d" in eventually_mono [OF lim])
       
  6404       apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+
       
  6405       done
       
  6406     have g_cint: "(\<lambda>u. g u/(u - w)) contour_integrable_on circlepath z r"
       
  6407       by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \<open>0 < r\<close>])
       
  6408     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"
       
  6409       by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \<open>0 < r\<close>])
       
  6410     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"
       
  6411       apply (rule Lim_transform_eventually [where f = "\<lambda>n. contour_integral (circlepath z r) (\<lambda>u. f n u/(u - w))"])
       
  6412       apply (rule eventually_mono [OF cont])
       
  6413       apply (rule contour_integral_unique [OF Cauchy_integral_circlepath])
       
  6414       using w
       
  6415       apply (auto simp: norm_minus_commute dist_norm cif_tends_cig)
       
  6416       done
       
  6417     have "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> 2 * of_real pi * \<i> * g w) F"
       
  6418       apply (rule tendsto_mult_left [OF tendstoI])
       
  6419       apply (rule eventually_mono [OF lim], assumption)
       
  6420       using w
       
  6421       apply (force simp add: dist_norm)
       
  6422       done
       
  6423     then have "((\<lambda>u. g u / (u - w)) has_contour_integral 2 * of_real pi * \<i> * g w) (circlepath z r)"
       
  6424       using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w
       
  6425       by (force simp add: dist_norm)
       
  6426     then have "((\<lambda>u. g u / (2 * of_real pi * \<i> * (u - w))) has_contour_integral g w) (circlepath z r)"
       
  6427       using has_contour_integral_div [where c = "2 * of_real pi * \<i>"]
       
  6428       by (force simp add: field_simps)
       
  6429     then show ?thesis
       
  6430       by (simp add: dist_norm)
       
  6431   qed
       
  6432   show ?thesis
       
  6433     using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified]
       
  6434     by (fastforce simp add: holomorphic_on_open contg intro: that)
       
  6435 qed
       
  6436 
       
  6437 
       
  6438 text\<open> Version showing that the limit is the limit of the derivatives.\<close>
       
  6439 
       
  6440 proposition has_complex_derivative_uniform_limit:
       
  6441   fixes z::complex
       
  6442   assumes cont: "eventually (\<lambda>n. continuous_on (cball z r) (f n) \<and>
       
  6443                                (\<forall>w \<in> ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F"
       
  6444       and lim: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> cball z r. norm(f n x - g x) < e) F"
       
  6445       and F:  "~ trivial_limit F" and "0 < r"
       
  6446   obtains g' where
       
  6447       "continuous_on (cball z r) g"
       
  6448       "\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F"
       
  6449 proof -
       
  6450   let ?conint = "contour_integral (circlepath z r)"
       
  6451   have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r"
       
  6452     by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] lim F];
       
  6453              auto simp: holomorphic_on_open field_differentiable_def)+
       
  6454   then obtain g' where g': "\<And>x. x \<in> ball z r \<Longrightarrow> (g has_field_derivative g' x) (at x)"
       
  6455     using DERIV_deriv_iff_has_field_derivative
       
  6456     by (fastforce simp add: holomorphic_on_open)
       
  6457   then have derg: "\<And>x. x \<in> ball z r \<Longrightarrow> deriv g x = g' x"
       
  6458     by (simp add: DERIV_imp_deriv)
       
  6459   have tends_f'n_g': "((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F" if w: "w \<in> ball z r" for w
       
  6460   proof -
       
  6461     have eq_f': "?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \<i>)"
       
  6462              if cont_fn: "continuous_on (cball z r) (f n)"
       
  6463              and fnd: "\<And>w. w \<in> ball z r \<Longrightarrow> (f n has_field_derivative f' n w) (at w)" for n
       
  6464     proof -
       
  6465       have hol_fn: "f n holomorphic_on ball z r"
       
  6466         using fnd by (force simp add: holomorphic_on_open)
       
  6467       have "(f n has_field_derivative 1 / (2 * of_real pi * \<i>) * ?conint (\<lambda>u. f n u / (u - w)\<^sup>2)) (at w)"
       
  6468         by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w])
       
  6469       then have f': "f' n w = 1 / (2 * of_real pi * \<i>) * ?conint (\<lambda>u. f n u / (u - w)\<^sup>2)"
       
  6470         using DERIV_unique [OF fnd] w by blast
       
  6471       show ?thesis
       
  6472         by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] divide_simps)
       
  6473     qed
       
  6474     define d where "d = (r - norm(w - z))^2"
       
  6475     have "d > 0"
       
  6476       using w by (simp add: dist_commute dist_norm d_def)
       
  6477     have dle: "\<And>y. r = cmod (z - y) \<Longrightarrow> d \<le> cmod ((y - w)\<^sup>2)"
       
  6478       apply (simp add: d_def norm_power)
       
  6479       apply (rule power_mono)
       
  6480       apply (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
       
  6481       apply (metis diff_ge_0_iff_ge dist_commute dist_norm less_eq_real_def mem_ball w)
       
  6482       done
       
  6483     have 1: "\<forall>\<^sub>F n in F. (\<lambda>x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r"
       
  6484       by (force simp add: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont])
       
  6485     have 2: "0 < e \<Longrightarrow> \<forall>\<^sub>F n in F. \<forall>x \<in> path_image (circlepath z r). cmod (f n x / (x - w)\<^sup>2 - g x / (x - w)\<^sup>2) < e" for e
       
  6486       using \<open>r > 0\<close>
       
  6487       apply (simp add: diff_divide_distrib [symmetric] norm_divide divide_simps sphere_def)
       
  6488       apply (rule eventually_mono [OF lim, of "e*d"])
       
  6489       apply (simp add: \<open>0 < d\<close>)
       
  6490       apply (force simp add: dist_norm dle intro: less_le_trans)
       
  6491       done
       
  6492     have "((\<lambda>n. contour_integral (circlepath z r) (\<lambda>x. f n x / (x - w)\<^sup>2))
       
  6493              \<longlongrightarrow> contour_integral (circlepath z r) ((\<lambda>x. g x / (x - w)\<^sup>2))) F"
       
  6494       by (rule Cauchy_Integral_Thm.contour_integral_uniform_limit_circlepath [OF 1 2 F \<open>0 < r\<close>])
       
  6495     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"
       
  6496       using Lim_null by (force intro!: tendsto_mult_right_zero)
       
  6497     have "((\<lambda>n. f' n w - g' w) \<longlongrightarrow> 0) F"
       
  6498       apply (rule Lim_transform_eventually [OF _ tendsto_0])
       
  6499       apply (force simp add: divide_simps intro: eq_f' eventually_mono [OF cont])
       
  6500       done
       
  6501     then show ?thesis using Lim_null by blast
       
  6502   qed
       
  6503   obtain g' where "\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F"
       
  6504       by (blast intro: tends_f'n_g' g' )
       
  6505   then show ?thesis using g
       
  6506     using that by blast
       
  6507 qed
       
  6508 
       
  6509 
       
  6510 subsection\<open>Some more simple/convenient versions for applications.\<close>
       
  6511 
       
  6512 lemma holomorphic_uniform_sequence:
       
  6513   assumes s: "open s"
       
  6514       and hol_fn: "\<And>n. (f n) holomorphic_on s"
       
  6515       and to_g: "\<And>x. x \<in> s
       
  6516                      \<Longrightarrow> \<exists>d. 0 < d \<and> cball x d \<subseteq> s \<and>
       
  6517                              (\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball x d. norm(f n y - g y) < e) sequentially)"
       
  6518   shows "g holomorphic_on s"
       
  6519 proof -
       
  6520   have "\<exists>f'. (g has_field_derivative f') (at z)" if "z \<in> s" for z
       
  6521   proof -
       
  6522     obtain r where "0 < r" and r: "cball z r \<subseteq> s"
       
  6523                and fg: "\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball z r. norm(f n y - g y) < e) sequentially"
       
  6524       using to_g [OF \<open>z \<in> s\<close>] by blast
       
  6525     have *: "\<forall>\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \<and> f n holomorphic_on ball z r"
       
  6526       apply (intro eventuallyI conjI)
       
  6527       using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r apply blast
       
  6528       apply (metis hol_fn holomorphic_on_subset interior_cball interior_subset r)
       
  6529       done
       
  6530     show ?thesis
       
  6531       apply (rule holomorphic_uniform_limit [OF *])
       
  6532       using \<open>0 < r\<close> centre_in_ball fg
       
  6533       apply (auto simp: holomorphic_on_open)
       
  6534       done
       
  6535   qed
       
  6536   with s show ?thesis
       
  6537     by (simp add: holomorphic_on_open)
       
  6538 qed
       
  6539 
       
  6540 lemma has_complex_derivative_uniform_sequence:
       
  6541   fixes s :: "complex set"
       
  6542   assumes s: "open s"
       
  6543       and hfd: "\<And>n x. x \<in> s \<Longrightarrow> ((f n) has_field_derivative f' n x) (at x)"
       
  6544       and to_g: "\<And>x. x \<in> s
       
  6545              \<Longrightarrow> \<exists>d. 0 < d \<and> cball x d \<subseteq> s \<and>
       
  6546                      (\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball x d. norm(f n y - g y) < e) sequentially)"
       
  6547   shows "\<exists>g'. \<forall>x \<in> s. (g has_field_derivative g' x) (at x) \<and> ((\<lambda>n. f' n x) \<longlongrightarrow> g' x) sequentially"
       
  6548 proof -
       
  6549   have y: "\<exists>y. (g has_field_derivative y) (at z) \<and> (\<lambda>n. f' n z) \<longlonglongrightarrow> y" if "z \<in> s" for z
       
  6550   proof -
       
  6551     obtain r where "0 < r" and r: "cball z r \<subseteq> s"
       
  6552                and fg: "\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball z r. norm(f n y - g y) < e) sequentially"
       
  6553       using to_g [OF \<open>z \<in> s\<close>] by blast
       
  6554     have *: "\<forall>\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \<and>
       
  6555                                    (\<forall>w \<in> ball z r. ((f n) has_field_derivative (f' n w)) (at w))"
       
  6556       apply (intro eventuallyI conjI)
       
  6557       apply (meson hfd holomorphic_on_imp_continuous_on holomorphic_on_open holomorphic_on_subset r s)
       
  6558       using ball_subset_cball hfd r apply blast
       
  6559       done
       
  6560     show ?thesis
       
  6561       apply (rule has_complex_derivative_uniform_limit [OF *, of g])
       
  6562       using \<open>0 < r\<close> centre_in_ball fg
       
  6563       apply force+
       
  6564       done
       
  6565   qed
       
  6566   show ?thesis
       
  6567     by (rule bchoice) (blast intro: y)
       
  6568 qed
       
  6569 
       
  6570 
       
  6571 subsection\<open>On analytic functions defined by a series.\<close>
       
  6572 
       
  6573 lemma series_and_derivative_comparison:
       
  6574   fixes s :: "complex set"
       
  6575   assumes s: "open s"
       
  6576       and h: "summable h"
       
  6577       and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
       
  6578       and to_g: "\<And>n x. \<lbrakk>N \<le> n; x \<in> s\<rbrakk> \<Longrightarrow> norm(f n x) \<le> h n"
       
  6579   obtains g g' where "\<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
       
  6580 proof -
       
  6581   obtain g where g: "\<And>e. e>0 \<Longrightarrow> \<exists>N. \<forall>n x. N \<le> n \<and> x \<in> s \<longrightarrow> dist (\<Sum>n<n. f n x) (g x) < e"
       
  6582     using series_comparison_uniform [OF h to_g, of N s] by force
       
  6583   have *: "\<exists>d>0. cball x d \<subseteq> s \<and> (\<forall>e>0. \<forall>\<^sub>F n in sequentially. \<forall>y\<in>cball x d. cmod ((\<Sum>a<n. f a y) - g y) < e)"
       
  6584          if "x \<in> s" for x
       
  6585   proof -
       
  6586     obtain d where "d>0" and d: "cball x d \<subseteq> s"
       
  6587       using open_contains_cball [of "s"] \<open>x \<in> s\<close> s by blast
       
  6588     then show ?thesis
       
  6589       apply (rule_tac x=d in exI)
       
  6590       apply (auto simp: dist_norm eventually_sequentially)
       
  6591       apply (metis g contra_subsetD dist_norm)
       
  6592       done
       
  6593   qed
       
  6594   have "(\<forall>x\<in>s. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> g x)"
       
  6595     using g by (force simp add: lim_sequentially)
       
  6596   moreover have "\<exists>g'. \<forall>x\<in>s. (g has_field_derivative g' x) (at x) \<and> (\<lambda>n. \<Sum>i<n. f' i x) \<longlonglongrightarrow> g' x"
       
  6597     by (rule has_complex_derivative_uniform_sequence [OF s]) (auto intro: * hfd DERIV_setsum)+
       
  6598   ultimately show ?thesis
       
  6599     by (force simp add: sums_def  conj_commute intro: that)
       
  6600 qed
       
  6601 
       
  6602 text\<open>A version where we only have local uniform/comparative convergence.\<close>
       
  6603 
       
  6604 lemma series_and_derivative_comparison_local:
       
  6605   fixes s :: "complex set"
       
  6606   assumes s: "open s"
       
  6607       and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
       
  6608       and to_g: "\<And>x. x \<in> s \<Longrightarrow>
       
  6609                       \<exists>d h N. 0 < d \<and> summable h \<and> (\<forall>n y. N \<le> n \<and> y \<in> ball x d \<longrightarrow> norm(f n y) \<le> h n)"
       
  6610   shows "\<exists>g g'. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
       
  6611 proof -
       
  6612   have "\<exists>y. (\<lambda>n. f n z) sums (\<Sum>n. f n z) \<and> (\<lambda>n. f' n z) sums y \<and> ((\<lambda>x. \<Sum>n. f n x) has_field_derivative y) (at z)"
       
  6613        if "z \<in> s" for z
       
  6614   proof -
       
  6615     obtain d h N where "0 < d" "summable h" and le_h: "\<And>n y. \<lbrakk>N \<le> n; y \<in> ball z d\<rbrakk> \<Longrightarrow> norm(f n y) \<le> h n"
       
  6616       using to_g \<open>z \<in> s\<close> by meson
       
  6617     then obtain r where "r>0" and r: "ball z r \<subseteq> ball z d \<inter> s" using \<open>z \<in> s\<close> s
       
  6618       by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq)
       
  6619     have 1: "open (ball z d \<inter> s)"
       
  6620       by (simp add: open_Int s)
       
  6621     have 2: "\<And>n x. x \<in> ball z d \<inter> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
       
  6622       by (auto simp: hfd)
       
  6623     obtain g g' where gg': "\<forall>x \<in> ball z d \<inter> s. ((\<lambda>n. f n x) sums g x) \<and>
       
  6624                                     ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
       
  6625       by (auto intro: le_h series_and_derivative_comparison [OF 1 \<open>summable h\<close> hfd])
       
  6626     then have "(\<lambda>n. f' n z) sums g' z"
       
  6627       by (meson \<open>0 < r\<close> centre_in_ball contra_subsetD r)
       
  6628     moreover have "(\<lambda>n. f n z) sums (\<Sum>n. f n z)"
       
  6629       by (metis summable_comparison_test' summable_sums centre_in_ball \<open>0 < d\<close> \<open>summable h\<close> le_h)
       
  6630     moreover have "((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' z) (at z)"
       
  6631       apply (rule_tac f=g in DERIV_transform_at [OF _ \<open>0 < r\<close>])
       
  6632       apply (simp add: gg' \<open>z \<in> s\<close> \<open>0 < d\<close>)
       
  6633       apply (metis (full_types) contra_subsetD dist_commute gg' mem_ball r sums_unique)
       
  6634       done
       
  6635     ultimately show ?thesis by auto
       
  6636   qed
       
  6637   then show ?thesis
       
  6638     by (rule_tac x="\<lambda>x. suminf  (\<lambda>n. f n x)" in exI) meson
       
  6639 qed
       
  6640 
       
  6641 
       
  6642 text\<open>Sometimes convenient to compare with a complex series of positive reals. (?)\<close>
       
  6643 
       
  6644 lemma series_and_derivative_comparison_complex:
       
  6645   fixes s :: "complex set"
       
  6646   assumes s: "open s"
       
  6647       and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
       
  6648       and to_g: "\<And>x. x \<in> s \<Longrightarrow>
       
  6649                       \<exists>d h N. 0 < d \<and> summable h \<and> range h \<subseteq> nonneg_Reals \<and> (\<forall>n y. N \<le> n \<and> y \<in> ball x d \<longrightarrow> cmod(f n y) \<le> cmod (h n))"
       
  6650   shows "\<exists>g g'. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
       
  6651 apply (rule series_and_derivative_comparison_local [OF s hfd], assumption)
       
  6652 apply (frule to_g)
       
  6653 apply (erule ex_forward)
       
  6654 apply (erule exE)
       
  6655 apply (rule_tac x="Re o h" in exI)
       
  6656 apply (erule ex_forward)
       
  6657 apply (simp add: summable_Re o_def )
       
  6658 apply (elim conjE all_forward)
       
  6659 apply (simp add: nonneg_Reals_cmod_eq_Re image_subset_iff)
       
  6660 done
       
  6661 
       
  6662 
       
  6663 text\<open>In particular, a power series is analytic inside circle of convergence.\<close>
       
  6664 
       
  6665 lemma power_series_and_derivative_0:
       
  6666   fixes a :: "nat \<Rightarrow> complex" and r::real
       
  6667   assumes "summable (\<lambda>n. a n * r^n)"
       
  6668     shows "\<exists>g g'. \<forall>z. cmod z < r \<longrightarrow>
       
  6669              ((\<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)"
       
  6670 proof (cases "0 < r")
       
  6671   case True
       
  6672     have der: "\<And>n z. ((\<lambda>x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)"
       
  6673       by (rule derivative_eq_intros | simp)+
       
  6674     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
       
  6675       using \<open>r > 0\<close>
       
  6676       apply (auto simp: algebra_simps norm_mult norm_divide norm_power of_real_add [symmetric] simp del: of_real_add)
       
  6677       using norm_triangle_ineq2 [of y z]
       
  6678       apply (simp only: diff_le_eq norm_minus_commute mult_2)
       
  6679       done
       
  6680     have "summable (\<lambda>n. a n * complex_of_real r ^ n)"
       
  6681       using assms \<open>r > 0\<close> by simp
       
  6682     moreover have "\<And>z. cmod z < r \<Longrightarrow> cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)"
       
  6683       using \<open>r > 0\<close>
       
  6684       by (simp add: of_real_add [symmetric] del: of_real_add)
       
  6685     ultimately have sum: "\<And>z. cmod z < r \<Longrightarrow> summable (\<lambda>n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)"
       
  6686       by (rule power_series_conv_imp_absconv_weak)
       
  6687     have "\<exists>g g'. \<forall>z \<in> ball 0 r. (\<lambda>n.  (a n) * z ^ n) sums g z \<and>
       
  6688                (\<lambda>n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \<and> (g has_field_derivative g' z) (at z)"
       
  6689       apply (rule series_and_derivative_comparison_complex [OF open_ball der])
       
  6690       apply (rule_tac x="(r - norm z)/2" in exI)
       
  6691       apply (simp add: dist_norm)
       
  6692       apply (rule_tac x="\<lambda>n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI)
       
  6693       using \<open>r > 0\<close>
       
  6694       apply (auto simp: sum nonneg_Reals_divide_I)
       
  6695       apply (rule_tac x=0 in exI)
       
  6696       apply (force simp: norm_mult norm_divide norm_power intro!: mult_left_mono power_mono y_le)
       
  6697       done
       
  6698   then show ?thesis
       
  6699     by (simp add: dist_0_norm ball_def)
       
  6700 next
       
  6701   case False then show ?thesis
       
  6702     apply (simp add: not_less)
       
  6703     using less_le_trans norm_not_less_zero by blast
       
  6704 qed
       
  6705 
       
  6706 proposition power_series_and_derivative:
       
  6707   fixes a :: "nat \<Rightarrow> complex" and r::real
       
  6708   assumes "summable (\<lambda>n. a n * r^n)"
       
  6709     obtains g g' where "\<forall>z \<in> ball w r.
       
  6710              ((\<lambda>n. a n * (z - w) ^ n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \<and>
       
  6711               (g has_field_derivative g' z) (at z)"
       
  6712   using power_series_and_derivative_0 [OF assms]
       
  6713   apply clarify
       
  6714   apply (rule_tac g="(\<lambda>z. g(z - w))" in that)
       
  6715   using DERIV_shift [where z="-w"]
       
  6716   apply (auto simp: norm_minus_commute Ball_def dist_norm)
       
  6717   done
       
  6718 
       
  6719 proposition power_series_holomorphic:
       
  6720   assumes "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>n. a n*(w - z)^n) sums f w)"
       
  6721     shows "f holomorphic_on ball z r"
       
  6722 proof -
       
  6723   have "\<exists>f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w
       
  6724   proof -
       
  6725     have inb: "z + complex_of_real ((dist z w + r) / 2) \<in> ball z r"
       
  6726     proof -
       
  6727       have wz: "cmod (w - z) < r" using w
       
  6728         by (auto simp: divide_simps dist_norm norm_minus_commute)
       
  6729       then have "0 \<le> r"
       
  6730         by (meson less_eq_real_def norm_ge_zero order_trans)
       
  6731       show ?thesis
       
  6732         using w by (simp add: dist_norm \<open>0\<le>r\<close> of_real_add [symmetric] del: of_real_add)
       
  6733     qed
       
  6734     have sum: "summable (\<lambda>n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))"
       
  6735       using assms [OF inb] by (force simp add: summable_def dist_norm)
       
  6736     obtain g g' where gg': "\<And>u. u \<in> ball z ((cmod (z - w) + r) / 2) \<Longrightarrow>
       
  6737                                (\<lambda>n. a n * (u - z) ^ n) sums g u \<and>
       
  6738                                (\<lambda>n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \<and> (g has_field_derivative g' u) (at u)"
       
  6739       by (rule power_series_and_derivative [OF sum, of z]) fastforce
       
  6740     have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u
       
  6741     proof -
       
  6742       have less: "cmod (z - u) * 2 < cmod (z - w) + r"
       
  6743         using that dist_triangle2 [of z u w]
       
  6744         by (simp add: dist_norm [symmetric] algebra_simps)
       
  6745       show ?thesis
       
  6746         apply (rule sums_unique2 [of "\<lambda>n. a n*(u - z)^n"])
       
  6747         using gg' [of u] less w
       
  6748         apply (auto simp: assms dist_norm)
       
  6749         done
       
  6750     qed
       
  6751     show ?thesis
       
  6752       apply (rule_tac x="g' w" in exI)
       
  6753       apply (rule DERIV_transform_at [where f=g and d="(r - norm(z - w))/2"])
       
  6754       using w gg' [of w]
       
  6755       apply (auto simp: dist_norm)
       
  6756       done
       
  6757   qed
       
  6758   then show ?thesis by (simp add: holomorphic_on_open)
       
  6759 qed
       
  6760 
       
  6761 corollary holomorphic_iff_power_series:
       
  6762      "f holomorphic_on ball z r \<longleftrightarrow>
       
  6763       (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
       
  6764   apply (intro iffI ballI)
       
  6765    using holomorphic_power_series  apply force
       
  6766   apply (rule power_series_holomorphic [where a = "\<lambda>n. (deriv ^^ n) f z / (fact n)"])
       
  6767   apply force
       
  6768   done
       
  6769 
       
  6770 corollary power_series_analytic:
       
  6771      "(\<And>w. w \<in> ball z r \<Longrightarrow> (\<lambda>n. a n*(w - z)^n) sums f w) \<Longrightarrow> f analytic_on ball z r"
       
  6772   by (force simp add: analytic_on_open intro!: power_series_holomorphic)
       
  6773 
       
  6774 corollary analytic_iff_power_series:
       
  6775      "f analytic_on ball z r \<longleftrightarrow>
       
  6776       (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
       
  6777   by (simp add: analytic_on_open holomorphic_iff_power_series)
       
  6778 
       
  6779 
       
  6780 subsection\<open>Equality between holomorphic functions, on open ball then connected set.\<close>
       
  6781 
       
  6782 lemma holomorphic_fun_eq_on_ball:
       
  6783    "\<lbrakk>f holomorphic_on ball z r; g holomorphic_on ball z r;
       
  6784      w \<in> ball z r;
       
  6785      \<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z\<rbrakk>
       
  6786      \<Longrightarrow> f w = g w"
       
  6787   apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
       
  6788   apply (auto simp: holomorphic_iff_power_series)
       
  6789   done
       
  6790 
       
  6791 lemma holomorphic_fun_eq_0_on_ball:
       
  6792    "\<lbrakk>f holomorphic_on ball z r;  w \<in> ball z r;
       
  6793      \<And>n. (deriv ^^ n) f z = 0\<rbrakk>
       
  6794      \<Longrightarrow> f w = 0"
       
  6795   apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
       
  6796   apply (auto simp: holomorphic_iff_power_series)
       
  6797   done
       
  6798 
       
  6799 lemma holomorphic_fun_eq_0_on_connected:
       
  6800   assumes holf: "f holomorphic_on s" and "open s"
       
  6801       and cons: "connected s"
       
  6802       and der: "\<And>n. (deriv ^^ n) f z = 0"
       
  6803       and "z \<in> s" "w \<in> s"
       
  6804     shows "f w = 0"
       
  6805 proof -
       
  6806   have *: "\<And>x e. \<lbrakk> \<forall>xa. (deriv ^^ xa) f x = 0;  ball x e \<subseteq> s\<rbrakk>
       
  6807            \<Longrightarrow> ball x e \<subseteq> (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
       
  6808     apply auto
       
  6809     apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv])
       
  6810     apply (rule holomorphic_on_subset [OF holf], simp_all)
       
  6811     by (metis funpow_add o_apply)
       
  6812   have 1: "openin (subtopology euclidean s) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
       
  6813     apply (rule open_subset, force)
       
  6814     using \<open>open s\<close>
       
  6815     apply (simp add: open_contains_ball Ball_def)
       
  6816     apply (erule all_forward)
       
  6817     using "*" by auto blast+
       
  6818   have 2: "closedin (subtopology euclidean s) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
       
  6819     using assms
       
  6820     by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv)
       
  6821   obtain e where "e>0" and e: "ball w e \<subseteq> s" using openE [OF \<open>open s\<close> \<open>w \<in> s\<close>] .
       
  6822   then have holfb: "f holomorphic_on ball w e"
       
  6823     using holf holomorphic_on_subset by blast
       
  6824   have 3: "(\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0}) = s \<Longrightarrow> f w = 0"
       
  6825     using \<open>e>0\<close> e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb])
       
  6826   show ?thesis
       
  6827     using cons der \<open>z \<in> s\<close>
       
  6828     apply (simp add: connected_clopen)
       
  6829     apply (drule_tac x="\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0}" in spec)
       
  6830     apply (auto simp: 1 2 3)
       
  6831     done
       
  6832 qed
       
  6833 
       
  6834 lemma holomorphic_fun_eq_on_connected:
       
  6835   assumes "f holomorphic_on s" "g holomorphic_on s" and "open s"  "connected s"
       
  6836       and "\<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z"
       
  6837       and "z \<in> s" "w \<in> s"
       
  6838     shows "f w = g w"
       
  6839   apply (rule holomorphic_fun_eq_0_on_connected [of "\<lambda>x. f x - g x" s z, simplified])
       
  6840   apply (intro assms holomorphic_intros)
       
  6841   using assms apply simp_all
       
  6842   apply (subst higher_deriv_diff, auto)
       
  6843   done
       
  6844 
       
  6845 lemma holomorphic_fun_eq_const_on_connected:
       
  6846   assumes holf: "f holomorphic_on s" and "open s"
       
  6847       and cons: "connected s"
       
  6848       and der: "\<And>n. 0 < n \<Longrightarrow> (deriv ^^ n) f z = 0"
       
  6849       and "z \<in> s" "w \<in> s"
       
  6850     shows "f w = f z"
       
  6851   apply (rule holomorphic_fun_eq_0_on_connected [of "\<lambda>w. f w - f z" s z, simplified])
       
  6852   apply (intro assms holomorphic_intros)
       
  6853   using assms apply simp_all
       
  6854   apply (subst higher_deriv_diff)
       
  6855   apply (intro holomorphic_intros | simp)+
       
  6856   done
       
  6857 
       
  6858 
       
  6859 subsection\<open>Some basic lemmas about poles/singularities.\<close>
       
  6860 
       
  6861 lemma pole_lemma:
       
  6862   assumes holf: "f holomorphic_on s" and a: "a \<in> interior s"
       
  6863     shows "(\<lambda>z. if z = a then deriv f a
       
  6864                  else (f z - f a) / (z - a)) holomorphic_on s" (is "?F holomorphic_on s")
       
  6865 proof -
       
  6866   have F1: "?F field_differentiable (at u within s)" if "u \<in> s" "u \<noteq> a" for u
       
  6867   proof -
       
  6868     have fcd: "f field_differentiable at u within s"
       
  6869       using holf holomorphic_on_def by (simp add: \<open>u \<in> s\<close>)
       
  6870     have cd: "(\<lambda>z. (f z - f a) / (z - a)) field_differentiable at u within s"
       
  6871       by (rule fcd derivative_intros | simp add: that)+
       
  6872     have "0 < dist a u" using that dist_nz by blast
       
  6873     then show ?thesis
       
  6874       by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \<open>u \<in> s\<close>)
       
  6875   qed
       
  6876   have F2: "?F field_differentiable at a" if "0 < e" "ball a e \<subseteq> s" for e
       
  6877   proof -
       
  6878     have holfb: "f holomorphic_on ball a e"
       
  6879       by (rule holomorphic_on_subset [OF holf \<open>ball a e \<subseteq> s\<close>])
       
  6880     have 2: "?F holomorphic_on ball a e - {a}"
       
  6881       apply (rule holomorphic_on_subset [where s = "s - {a}"])
       
  6882       apply (simp add: holomorphic_on_def field_differentiable_def [symmetric])
       
  6883       using mem_ball that
       
  6884       apply (auto intro: F1 field_differentiable_within_subset)
       
  6885       done
       
  6886     have "isCont (\<lambda>z. if z = a then deriv f a else (f z - f a) / (z - a)) x"
       
  6887             if "dist a x < e" for x
       
  6888     proof (cases "x=a")
       
  6889       case True then show ?thesis
       
  6890       using holfb \<open>0 < e\<close>
       
  6891       apply (simp add: holomorphic_on_open field_differentiable_def [symmetric])
       
  6892       apply (drule_tac x=a in bspec)
       
  6893       apply (auto simp: DERIV_deriv_iff_field_differentiable [symmetric] continuous_at DERIV_iff2
       
  6894                 elim: rev_iffD1 [OF _ LIM_equal])
       
  6895       done
       
  6896     next
       
  6897       case False with 2 that show ?thesis
       
  6898         by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at)
       
  6899     qed
       
  6900     then have 1: "continuous_on (ball a e) ?F"
       
  6901       by (clarsimp simp:  continuous_on_eq_continuous_at)
       
  6902     have "?F holomorphic_on ball a e"
       
  6903       by (auto intro: no_isolated_singularity [OF 1 2])
       
  6904     with that show ?thesis
       
  6905       by (simp add: holomorphic_on_open field_differentiable_def [symmetric]
       
  6906                     field_differentiable_at_within)
       
  6907   qed
       
  6908   show ?thesis
       
  6909   proof
       
  6910     fix x assume "x \<in> s" show "?F field_differentiable at x within s"
       
  6911     proof (cases "x=a")
       
  6912       case True then show ?thesis
       
  6913       using a by (auto simp: mem_interior intro: field_differentiable_at_within F2)
       
  6914     next
       
  6915       case False with F1 \<open>x \<in> s\<close>
       
  6916       show ?thesis by blast
       
  6917     qed
       
  6918   qed
       
  6919 qed
       
  6920 
       
  6921 proposition pole_theorem:
       
  6922   assumes holg: "g holomorphic_on s" and a: "a \<in> interior s"
       
  6923       and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
       
  6924     shows "(\<lambda>z. if z = a then deriv g a
       
  6925                  else f z - g a/(z - a)) holomorphic_on s"
       
  6926   using pole_lemma [OF holg a]
       
  6927   by (rule holomorphic_transform) (simp add: eq divide_simps)
       
  6928 
       
  6929 lemma pole_lemma_open:
       
  6930   assumes "f holomorphic_on s" "open s"
       
  6931     shows "(\<lambda>z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on s"
       
  6932 proof (cases "a \<in> s")
       
  6933   case True with assms interior_eq pole_lemma
       
  6934     show ?thesis by fastforce
       
  6935 next
       
  6936   case False with assms show ?thesis
       
  6937     apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify)
       
  6938     apply (rule field_differentiable_transform_within [where f = "\<lambda>z. (f z - f a)/(z - a)" and d = 1])
       
  6939     apply (rule derivative_intros | force)+
       
  6940     done
       
  6941 qed
       
  6942 
       
  6943 proposition pole_theorem_open:
       
  6944   assumes holg: "g holomorphic_on s" and s: "open s"
       
  6945       and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
       
  6946     shows "(\<lambda>z. if z = a then deriv g a
       
  6947                  else f z - g a/(z - a)) holomorphic_on s"
       
  6948   using pole_lemma_open [OF holg s]
       
  6949   by (rule holomorphic_transform) (auto simp: eq divide_simps)
       
  6950 
       
  6951 proposition pole_theorem_0:
       
  6952   assumes holg: "g holomorphic_on s" and a: "a \<in> interior s"
       
  6953       and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
       
  6954       and [simp]: "f a = deriv g a" "g a = 0"
       
  6955     shows "f holomorphic_on s"
       
  6956   using pole_theorem [OF holg a eq]
       
  6957   by (rule holomorphic_transform) (auto simp: eq divide_simps)
       
  6958 
       
  6959 proposition pole_theorem_open_0:
       
  6960   assumes holg: "g holomorphic_on s" and s: "open s"
       
  6961       and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
       
  6962       and [simp]: "f a = deriv g a" "g a = 0"
       
  6963     shows "f holomorphic_on s"
       
  6964   using pole_theorem_open [OF holg s eq]
       
  6965   by (rule holomorphic_transform) (auto simp: eq divide_simps)
       
  6966 
       
  6967 lemma pole_theorem_analytic:
       
  6968   assumes g: "g analytic_on s"
       
  6969       and eq: "\<And>z. z \<in> s
       
  6970              \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>w \<in> ball z d - {a}. g w = (w - a) * f w)"
       
  6971     shows "(\<lambda>z. if z = a then deriv g a
       
  6972                  else f z - g a/(z - a)) analytic_on s"
       
  6973 using g
       
  6974 apply (simp add: analytic_on_def Ball_def)
       
  6975 apply (safe elim!: all_forward dest!: eq)
       
  6976 apply (rule_tac x="min d e" in exI, simp)
       
  6977 apply (rule pole_theorem_open)
       
  6978 apply (auto simp: holomorphic_on_subset subset_ball)
       
  6979 done
       
  6980 
       
  6981 lemma pole_theorem_analytic_0:
       
  6982   assumes g: "g analytic_on s"
       
  6983       and eq: "\<And>z. z \<in> s \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>w \<in> ball z d - {a}. g w = (w - a) * f w)"
       
  6984       and [simp]: "f a = deriv g a" "g a = 0"
       
  6985     shows "f analytic_on s"
       
  6986 proof -
       
  6987   have [simp]: "(\<lambda>z. if z = a then deriv g a else f z - g a / (z - a)) = f"
       
  6988     by auto
       
  6989   show ?thesis
       
  6990     using pole_theorem_analytic [OF g eq] by simp
       
  6991 qed
       
  6992 
       
  6993 lemma pole_theorem_analytic_open_superset:
       
  6994   assumes g: "g analytic_on s" and "s \<subseteq> t" "open t"
       
  6995       and eq: "\<And>z. z \<in> t - {a} \<Longrightarrow> g z = (z - a) * f z"
       
  6996     shows "(\<lambda>z. if z = a then deriv g a
       
  6997                  else f z - g a/(z - a)) analytic_on s"
       
  6998   apply (rule pole_theorem_analytic [OF g])
       
  6999   apply (rule openE [OF \<open>open t\<close>])
       
  7000   using assms eq by auto
       
  7001 
       
  7002 lemma pole_theorem_analytic_open_superset_0:
       
  7003   assumes g: "g analytic_on s" "s \<subseteq> t" "open t" "\<And>z. z \<in> t - {a} \<Longrightarrow> g z = (z - a) * f z"
       
  7004       and [simp]: "f a = deriv g a" "g a = 0"
       
  7005     shows "f analytic_on s"
       
  7006 proof -
       
  7007   have [simp]: "(\<lambda>z. if z = a then deriv g a else f z - g a / (z - a)) = f"
       
  7008     by auto
       
  7009   have "(\<lambda>z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on s"
       
  7010     by (rule pole_theorem_analytic_open_superset [OF g])
       
  7011   then show ?thesis by simp
       
  7012 qed
       
  7013 
       
  7014 
       
  7015 
       
  7016 subsection\<open>General, homology form of Cauchy's theorem.\<close>
       
  7017 
       
  7018 text\<open>Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\<close>
       
  7019 
       
  7020 lemma contour_integral_continuous_on_linepath_2D:
       
  7021   assumes "open u" and cont_dw: "\<And>w. w \<in> u \<Longrightarrow> F w contour_integrable_on (linepath a b)"
       
  7022       and cond_uu: "continuous_on (u \<times> u) (\<lambda>(x,y). F x y)"
       
  7023       and abu: "closed_segment a b \<subseteq> u"
       
  7024     shows "continuous_on u (\<lambda>w. contour_integral (linepath a b) (F w))"
       
  7025 proof -
       
  7026   have *: "\<exists>d>0. \<forall>x'\<in>u. dist x' w < d \<longrightarrow>
       
  7027                          dist (contour_integral (linepath a b) (F x'))
       
  7028                               (contour_integral (linepath a b) (F w)) \<le> \<epsilon>"
       
  7029           if "w \<in> u" "0 < \<epsilon>" "a \<noteq> b" for w \<epsilon>
       
  7030   proof -
       
  7031     obtain \<delta> where "\<delta>>0" and \<delta>: "cball w \<delta> \<subseteq> u" using open_contains_cball \<open>open u\<close> \<open>w \<in> u\<close> by force
       
  7032     let ?TZ = "{(t,z) |t z. t \<in> cball w \<delta> \<and> z \<in> closed_segment a b}"
       
  7033     have "uniformly_continuous_on ?TZ (\<lambda>(x,y). F x y)"
       
  7034       apply (rule compact_uniformly_continuous)
       
  7035       apply (rule continuous_on_subset[OF cond_uu])
       
  7036       using abu \<delta>
       
  7037       apply (auto simp: compact_Times simp del: mem_cball)
       
  7038       done
       
  7039     then obtain \<eta> where "\<eta>>0"
       
  7040         and \<eta>: "\<And>x x'. \<lbrakk>x\<in>?TZ; x'\<in>?TZ; dist x' x < \<eta>\<rbrakk> \<Longrightarrow>
       
  7041                          dist ((\<lambda>(x,y). F x y) x') ((\<lambda>(x,y). F x y) x) < \<epsilon>/norm(b - a)"
       
  7042       apply (rule uniformly_continuous_onE [where e = "\<epsilon>/norm(b - a)"])
       
  7043       using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> by auto
       
  7044     have \<eta>: "\<lbrakk>norm (w - x1) \<le> \<delta>;   x2 \<in> closed_segment a b;
       
  7045               norm (w - x1') \<le> \<delta>;  x2' \<in> closed_segment a b; norm ((x1', x2') - (x1, x2)) < \<eta>\<rbrakk>
       
  7046               \<Longrightarrow> norm (F x1' x2' - F x1 x2) \<le> \<epsilon> / cmod (b - a)"
       
  7047              for x1 x2 x1' x2'
       
  7048       using \<eta> [of "(x1,x2)" "(x1',x2')"] by (force simp add: dist_norm)
       
  7049     have le_ee: "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>"
       
  7050                 if "x' \<in> u" "cmod (x' - w) < \<delta>" "cmod (x' - w) < \<eta>"  for x'
       
  7051     proof -
       
  7052       have "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>/norm(b - a) * norm(b - a)"
       
  7053         apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \<eta>])
       
  7054         apply (rule contour_integrable_diff [OF cont_dw cont_dw])
       
  7055         using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> \<open>0 < \<delta>\<close> \<open>w \<in> u\<close> that
       
  7056         apply (auto simp: norm_minus_commute)
       
  7057         done
       
  7058       also have "... = \<epsilon>" using \<open>a \<noteq> b\<close> by simp
       
  7059       finally show ?thesis .
       
  7060     qed
       
  7061     show ?thesis
       
  7062       apply (rule_tac x="min \<delta> \<eta>" in exI)
       
  7063       using \<open>0 < \<delta>\<close> \<open>0 < \<eta>\<close>
       
  7064       apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \<open>w \<in> u\<close> intro: le_ee)
       
  7065       done
       
  7066   qed
       
  7067   show ?thesis
       
  7068     apply (rule continuous_onI)
       
  7069     apply (cases "a=b")
       
  7070     apply (auto intro: *)
       
  7071     done
       
  7072 qed
       
  7073 
       
  7074 text\<open>This version has @{term"polynomial_function \<gamma>"} as an additional assumption.\<close>
       
  7075 lemma Cauchy_integral_formula_global_weak:
       
  7076     assumes u: "open u" and holf: "f holomorphic_on u"
       
  7077         and z: "z \<in> u" and \<gamma>: "polynomial_function \<gamma>"
       
  7078         and pasz: "path_image \<gamma> \<subseteq> u - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  7079         and zero: "\<And>w. w \<notin> u \<Longrightarrow> winding_number \<gamma> w = 0"
       
  7080       shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
       
  7081 proof -
       
  7082   obtain \<gamma>' where pf\<gamma>': "polynomial_function \<gamma>'" and \<gamma>': "\<And>x. (\<gamma> has_vector_derivative (\<gamma>' x)) (at x)"
       
  7083     using has_vector_derivative_polynomial_function [OF \<gamma>] by blast
       
  7084   then have "bounded(path_image \<gamma>')"
       
  7085     by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function)
       
  7086   then obtain B where "B>0" and B: "\<And>x. x \<in> path_image \<gamma>' \<Longrightarrow> norm x \<le> B"
       
  7087     using bounded_pos by force
       
  7088   define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w
       
  7089   define v where "v = {w. w \<notin> path_image \<gamma> \<and> winding_number \<gamma> w = 0}"
       
  7090   have "path \<gamma>" "valid_path \<gamma>" using \<gamma>
       
  7091     by (auto simp: path_polynomial_function valid_path_polynomial_function)
       
  7092   then have ov: "open v"
       
  7093     by (simp add: v_def open_winding_number_levelsets loop)
       
  7094   have uv_Un: "u \<union> v = UNIV"
       
  7095     using pasz zero by (auto simp: v_def)
       
  7096   have conf: "continuous_on u f"
       
  7097     by (metis holf holomorphic_on_imp_continuous_on)
       
  7098   have hol_d: "(d y) holomorphic_on u" if "y \<in> u" for y
       
  7099   proof -
       
  7100     have *: "(\<lambda>c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on u"
       
  7101       by (simp add: holf pole_lemma_open u)
       
  7102     then have "isCont (\<lambda>x. if x = y then deriv f y else (f x - f y) / (x - y)) y"
       
  7103       using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that u by fastforce
       
  7104     then have "continuous_on u (d y)"
       
  7105       apply (simp add: d_def continuous_on_eq_continuous_at u, clarify)
       
  7106       using * holomorphic_on_def
       
  7107       by (meson field_differentiable_within_open field_differentiable_imp_continuous_at u)
       
  7108     moreover have "d y holomorphic_on u - {y}"
       
  7109       apply (simp add: d_def holomorphic_on_open u open_delete field_differentiable_def [symmetric])
       
  7110       apply (intro ballI)
       
  7111       apply (rename_tac w)
       
  7112       apply (rule_tac d="dist w y" and f = "\<lambda>w. (f w - f y)/(w - y)" in field_differentiable_transform_within)
       
  7113       apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros)
       
  7114       using analytic_on_imp_differentiable_at analytic_on_open holf u apply blast
       
  7115       done
       
  7116     ultimately show ?thesis
       
  7117       by (rule no_isolated_singularity) (auto simp: u)
       
  7118   qed
       
  7119   have cint_fxy: "(\<lambda>x. (f x - f y) / (x - y)) contour_integrable_on \<gamma>" if "y \<notin> path_image \<gamma>" for y
       
  7120     apply (rule contour_integrable_holomorphic_simple [where s = "u-{y}"])
       
  7121     using \<open>valid_path \<gamma>\<close> pasz
       
  7122     apply (auto simp: u open_delete)
       
  7123     apply (rule continuous_intros holomorphic_intros continuous_on_subset [OF conf] holomorphic_on_subset [OF holf] |
       
  7124                 force simp add: that)+
       
  7125     done
       
  7126   define h where
       
  7127     "h z = (if z \<in> u then contour_integral \<gamma> (d z) else contour_integral \<gamma> (\<lambda>w. f w/(w - z)))" for z
       
  7128   have U: "\<And>z. z \<in> u \<Longrightarrow> ((d z) has_contour_integral h z) \<gamma>"
       
  7129     apply (simp add: h_def)
       
  7130     apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]])
       
  7131     using u pasz \<open>valid_path \<gamma>\<close>
       
  7132     apply (auto intro: holomorphic_on_imp_continuous_on hol_d)
       
  7133     done
       
  7134   have V: "((\<lambda>w. f w / (w - z)) has_contour_integral h z) \<gamma>" if z: "z \<in> v" for z
       
  7135   proof -
       
  7136     have 0: "0 = (f z) * 2 * of_real (2 * pi) * \<i> * winding_number \<gamma> z"
       
  7137       using v_def z by auto
       
  7138     then have "((\<lambda>x. 1 / (x - z)) has_contour_integral 0) \<gamma>"
       
  7139      using z v_def  has_contour_integral_winding_number [OF \<open>valid_path \<gamma>\<close>] by fastforce
       
  7140     then have "((\<lambda>x. f z * (1 / (x - z))) has_contour_integral 0) \<gamma>"
       
  7141       using has_contour_integral_lmul by fastforce
       
  7142     then have "((\<lambda>x. f z / (x - z)) has_contour_integral 0) \<gamma>"
       
  7143       by (simp add: divide_simps)
       
  7144     moreover have "((\<lambda>x. (f x - f z) / (x - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
       
  7145       using z
       
  7146       apply (auto simp: v_def)
       
  7147       apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy)
       
  7148       done
       
  7149     ultimately have *: "((\<lambda>x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \<gamma> (d z))) \<gamma>"
       
  7150       by (rule has_contour_integral_add)
       
  7151     have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
       
  7152             if  "z \<in> u"
       
  7153       using * by (auto simp: divide_simps has_contour_integral_eq)
       
  7154     moreover have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (\<lambda>w. f w / (w - z))) \<gamma>"
       
  7155             if "z \<notin> u"
       
  7156       apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]])
       
  7157       using u pasz \<open>valid_path \<gamma>\<close> that
       
  7158       apply (auto intro: holomorphic_on_imp_continuous_on hol_d)
       
  7159       apply (rule continuous_intros conf holomorphic_intros holf | force)+
       
  7160       done
       
  7161     ultimately show ?thesis
       
  7162       using z by (simp add: h_def)
       
  7163   qed
       
  7164   have znot: "z \<notin> path_image \<gamma>"
       
  7165     using pasz by blast
       
  7166   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"
       
  7167     using separate_compact_closed [of "path_image \<gamma>" "-u"] pasz u
       
  7168     by (fastforce simp add: \<open>path \<gamma>\<close> compact_path_image)
       
  7169   obtain dd where "0 < dd" and dd: "{y + k | y k. y \<in> path_image \<gamma> \<and> k \<in> ball 0 dd} \<subseteq> u"
       
  7170     apply (rule that [of "d0/2"])
       
  7171     using \<open>0 < d0\<close>
       
  7172     apply (auto simp: dist_norm dest: d0)
       
  7173     done
       
  7174   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"
       
  7175     apply (rule_tac x=x in exI)
       
  7176     apply (rule_tac x="x'-x" in exI)
       
  7177     apply (force simp add: dist_norm)
       
  7178     done
       
  7179   then have 1: "path_image \<gamma> \<subseteq> interior {y + k |y k. y \<in> path_image \<gamma> \<and> k \<in> cball 0 (dd / 2)}"
       
  7180     apply (clarsimp simp add: mem_interior)
       
  7181     using \<open>0 < dd\<close>
       
  7182     apply (rule_tac x="dd/2" in exI, auto)
       
  7183     done
       
  7184   obtain t where "compact t" and subt: "path_image \<gamma> \<subseteq> interior t" and t: "t \<subseteq> u"
       
  7185     apply (rule that [OF _ 1])
       
  7186     apply (fastforce simp add: \<open>valid_path \<gamma>\<close> compact_valid_path_image intro!: compact_sums)
       
  7187     apply (rule order_trans [OF _ dd])
       
  7188     using \<open>0 < dd\<close> by fastforce
       
  7189   obtain L where "L>0"
       
  7190            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>
       
  7191                          cmod (contour_integral \<gamma> f) \<le> L * B"
       
  7192       using contour_integral_bound_exists [OF open_interior \<open>valid_path \<gamma>\<close> subt]
       
  7193       by blast
       
  7194   have "bounded(f ` t)"
       
  7195     by (meson \<open>compact t\<close> compact_continuous_image compact_imp_bounded conf continuous_on_subset t)
       
  7196   then obtain D where "D>0" and D: "\<And>x. x \<in> t \<Longrightarrow> norm (f x) \<le> D"
       
  7197     by (auto simp: bounded_pos)
       
  7198   obtain C where "C>0" and C: "\<And>x. x \<in> t \<Longrightarrow> norm x \<le> C"
       
  7199     using \<open>compact t\<close> bounded_pos compact_imp_bounded by force
       
  7200   have "dist (h y) 0 \<le> e" if "0 < e" and le: "D * L / e + C \<le> cmod y" for e y
       
  7201   proof -
       
  7202     have "D * L / e > 0"  using \<open>D>0\<close> \<open>L>0\<close> \<open>e>0\<close> by simp
       
  7203     with le have ybig: "norm y > C" by force
       
  7204     with C have "y \<notin> t"  by force
       
  7205     then have ynot: "y \<notin> path_image \<gamma>"
       
  7206       using subt interior_subset by blast
       
  7207     have [simp]: "winding_number \<gamma> y = 0"
       
  7208       apply (rule winding_number_zero_outside [of _ "cball 0 C"])
       
  7209       using ybig interior_subset subt
       
  7210       apply (force simp add: loop \<open>path \<gamma>\<close> dist_norm intro!: C)+
       
  7211       done
       
  7212     have [simp]: "h y = contour_integral \<gamma> (\<lambda>w. f w/(w - y))"
       
  7213       by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V)
       
  7214     have holint: "(\<lambda>w. f w / (w - y)) holomorphic_on interior t"
       
  7215       apply (rule holomorphic_on_divide)
       
  7216       using holf holomorphic_on_subset interior_subset t apply blast
       
  7217       apply (rule holomorphic_intros)+
       
  7218       using \<open>y \<notin> t\<close> interior_subset by auto
       
  7219     have leD: "cmod (f z / (z - y)) \<le> D * (e / L / D)" if z: "z \<in> interior t" for z
       
  7220     proof -
       
  7221       have "D * L / e + cmod z \<le> cmod y"
       
  7222         using le C [of z] z using interior_subset by force
       
  7223       then have DL2: "D * L / e \<le> cmod (z - y)"
       
  7224         using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute)
       
  7225       have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))"
       
  7226         by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse)
       
  7227       also have "... \<le> D * (e / L / D)"
       
  7228         apply (rule mult_mono)
       
  7229         using that D interior_subset apply blast
       
  7230         using \<open>L>0\<close> \<open>e>0\<close> \<open>D>0\<close> DL2
       
  7231         apply (auto simp: norm_divide divide_simps algebra_simps)
       
  7232         done
       
  7233       finally show ?thesis .
       
  7234     qed
       
  7235     have "dist (h y) 0 = cmod (contour_integral \<gamma> (\<lambda>w. f w / (w - y)))"
       
  7236       by (simp add: dist_norm)
       
  7237     also have "... \<le> L * (D * (e / L / D))"
       
  7238       by (rule L [OF holint leD])
       
  7239     also have "... = e"
       
  7240       using  \<open>L>0\<close> \<open>0 < D\<close> by auto
       
  7241     finally show ?thesis .
       
  7242   qed
       
  7243   then have "(h \<longlongrightarrow> 0) at_infinity"
       
  7244     by (meson Lim_at_infinityI)
       
  7245   moreover have "h holomorphic_on UNIV"
       
  7246   proof -
       
  7247     have con_ff: "continuous (at (x,z)) (\<lambda>(x,y). (f y - f x) / (y - x))"
       
  7248                  if "x \<in> u" "z \<in> u" "x \<noteq> z" for x z
       
  7249       using that conf
       
  7250       apply (simp add: split_def continuous_on_eq_continuous_at u)
       
  7251       apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+
       
  7252       done
       
  7253     have con_fstsnd: "continuous_on UNIV (\<lambda>x. (fst x - snd x) ::complex)"
       
  7254       by (rule continuous_intros)+
       
  7255     have open_uu_Id: "open (u \<times> u - Id)"
       
  7256       apply (rule open_Diff)
       
  7257       apply (simp add: open_Times u)
       
  7258       using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0]
       
  7259       apply (auto simp: Id_fstsnd_eq algebra_simps)
       
  7260       done
       
  7261     have con_derf: "continuous (at z) (deriv f)" if "z \<in> u" for z
       
  7262       apply (rule continuous_on_interior [of u])
       
  7263       apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on u)
       
  7264       by (simp add: interior_open that u)
       
  7265     have tendsto_f': "((\<lambda>(x,y). if y = x then deriv f (x)
       
  7266                                 else (f (y) - f (x)) / (y - x)) \<longlongrightarrow> deriv f x)
       
  7267                       (at (x, x) within u \<times> u)" if "x \<in> u" for x
       
  7268     proof (rule Lim_withinI)
       
  7269       fix e::real assume "0 < e"
       
  7270       obtain k1 where "k1>0" and k1: "\<And>x'. norm (x' - x) \<le> k1 \<Longrightarrow> norm (deriv f x' - deriv f x) < e"
       
  7271         using \<open>0 < e\<close> continuous_within_E [OF con_derf [OF \<open>x \<in> u\<close>]]
       
  7272         by (metis UNIV_I dist_norm)
       
  7273       obtain k2 where "k2>0" and k2: "ball x k2 \<subseteq> u" by (blast intro: openE [OF u] \<open>x \<in> u\<close>)
       
  7274       have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e"
       
  7275                     if "z' \<noteq> x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2"
       
  7276                  for x' z'
       
  7277       proof -
       
  7278         have cs_less: "w \<in> closed_segment x' z' \<Longrightarrow> cmod (w - x) \<le> norm (x'-x, z'-x)" for w
       
  7279           apply (drule segment_furthest_le [where y=x])
       
  7280           by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans)
       
  7281         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
       
  7282           by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans)
       
  7283         have f_has_der: "\<And>x. x \<in> u \<Longrightarrow> (f has_field_derivative deriv f x) (at x within u)"
       
  7284           by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def u)
       
  7285         have "closed_segment x' z' \<subseteq> u"
       
  7286           by (rule order_trans [OF _ k2]) (simp add: cs_less  le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff)
       
  7287         then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')"
       
  7288           using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz  by simp
       
  7289         then have *: "((\<lambda>x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')"
       
  7290           by (rule has_contour_integral_div)
       
  7291         have "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e/norm(z' - x') * norm(z' - x')"
       
  7292           apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]])
       
  7293           using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']]
       
  7294                  \<open>e > 0\<close>  \<open>z' \<noteq> x'\<close>
       
  7295           apply (auto simp: norm_divide divide_simps derf_le)
       
  7296           done
       
  7297         also have "... \<le> e" using \<open>0 < e\<close> by simp
       
  7298         finally show ?thesis .
       
  7299       qed
       
  7300       show "\<exists>d>0. \<forall>xa\<in>u \<times> u.
       
  7301                   0 < dist xa (x, x) \<and> dist xa (x, x) < d \<longrightarrow>
       
  7302                   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"
       
  7303         apply (rule_tac x="min k1 k2" in exI)
       
  7304         using \<open>k1>0\<close> \<open>k2>0\<close> \<open>e>0\<close>
       
  7305         apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le)
       
  7306         done
       
  7307     qed
       
  7308     have con_pa_f: "continuous_on (path_image \<gamma>) f"
       
  7309       by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt t)
       
  7310     have le_B: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (vector_derivative \<gamma> (at t)) \<le> B"
       
  7311       apply (rule B)
       
  7312       using \<gamma>' using path_image_def vector_derivative_at by fastforce
       
  7313     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>"
       
  7314       by (simp add: V)
       
  7315     have cond_uu: "continuous_on (u \<times> u) (\<lambda>(x,y). d x y)"
       
  7316       apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f')
       
  7317       apply (simp add: tendsto_within_open_NO_MATCH open_Times u, clarify)
       
  7318       apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\<lambda>(x,y). (f y - f x) / (y - x))"])
       
  7319       using con_ff
       
  7320       apply (auto simp: continuous_within)
       
  7321       done
       
  7322     have hol_dw: "(\<lambda>z. d z w) holomorphic_on u" if "w \<in> u" for w
       
  7323     proof -
       
  7324       have "continuous_on u ((\<lambda>(x,y). d x y) o (\<lambda>z. (w,z)))"
       
  7325         by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+
       
  7326       then have *: "continuous_on u (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z))"
       
  7327         by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps)
       
  7328       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"
       
  7329         apply (rule_tac f = "\<lambda>x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within)
       
  7330         apply (rule u derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp add: dist_commute)+
       
  7331         done
       
  7332       show ?thesis
       
  7333         unfolding d_def
       
  7334         apply (rule no_isolated_singularity [OF * _ u, where k = "{w}"])
       
  7335         apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff u **)
       
  7336         done
       
  7337     qed
       
  7338     { fix a b
       
  7339       assume abu: "closed_segment a b \<subseteq> u"
       
  7340       then have "\<And>w. w \<in> u \<Longrightarrow> (\<lambda>z. d z w) contour_integrable_on (linepath a b)"
       
  7341         by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on)
       
  7342       then have cont_cint_d: "continuous_on u (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
       
  7343         apply (rule contour_integral_continuous_on_linepath_2D [OF \<open>open u\<close> _ _ abu])
       
  7344         apply (auto simp: intro: continuous_on_swap_args cond_uu)
       
  7345         done
       
  7346       have cont_cint_d\<gamma>: "continuous_on {0..1} ((\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) o \<gamma>)"
       
  7347         apply (rule continuous_on_compose)
       
  7348         using \<open>path \<gamma>\<close> path_def pasz
       
  7349         apply (auto intro!: continuous_on_subset [OF cont_cint_d])
       
  7350         apply (force simp add: path_image_def)
       
  7351         done
       
  7352       have cint_cint: "(\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) contour_integrable_on \<gamma>"
       
  7353         apply (simp add: contour_integrable_on)
       
  7354         apply (rule integrable_continuous_real)
       
  7355         apply (rule continuous_on_mult [OF cont_cint_d\<gamma> [unfolded o_def]])
       
  7356         using pf\<gamma>'
       
  7357         by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \<gamma>'])
       
  7358       have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\<lambda>z. contour_integral \<gamma> (d z))"
       
  7359         using abu  by (force simp add: h_def intro: contour_integral_eq)
       
  7360       also have "... =  contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
       
  7361         apply (rule contour_integral_swap)
       
  7362         apply (rule continuous_on_subset [OF cond_uu])
       
  7363         using abu pasz \<open>valid_path \<gamma>\<close>
       
  7364         apply (auto intro!: continuous_intros)
       
  7365         by (metis \<gamma>' continuous_on_eq path_def path_polynomial_function pf\<gamma>' vector_derivative_at)
       
  7366       finally have cint_h_eq:
       
  7367           "contour_integral (linepath a b) h =
       
  7368                     contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))" .
       
  7369       note cint_cint cint_h_eq
       
  7370     } note cint_h = this
       
  7371     have conthu: "continuous_on u h"
       
  7372     proof (simp add: continuous_on_sequentially, clarify)
       
  7373       fix a x
       
  7374       assume x: "x \<in> u" and au: "\<forall>n. a n \<in> u" and ax: "a \<longlonglongrightarrow> x"
       
  7375       then have A1: "\<forall>\<^sub>F n in sequentially. d (a n) contour_integrable_on \<gamma>"
       
  7376         by (meson U contour_integrable_on_def eventuallyI)
       
  7377       obtain dd where "dd>0" and dd: "cball x dd \<subseteq> u" using open_contains_cball u x by force
       
  7378       have A2: "\<forall>\<^sub>F n in sequentially. \<forall>xa\<in>path_image \<gamma>. cmod (d (a n) xa - d x xa) < ee" if "0 < ee" for ee
       
  7379       proof -
       
  7380         let ?ddpa = "{(w,z) |w z. w \<in> cball x dd \<and> z \<in> path_image \<gamma>}"
       
  7381         have "uniformly_continuous_on ?ddpa (\<lambda>(x,y). d x y)"
       
  7382           apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]])
       
  7383           using dd pasz \<open>valid_path \<gamma>\<close>
       
  7384           apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball)
       
  7385           done
       
  7386         then obtain kk where "kk>0"
       
  7387             and kk: "\<And>x x'. \<lbrakk>x\<in>?ddpa; x'\<in>?ddpa; dist x' x < kk\<rbrakk> \<Longrightarrow>
       
  7388                              dist ((\<lambda>(x,y). d x y) x') ((\<lambda>(x,y). d x y) x) < ee"
       
  7389           apply (rule uniformly_continuous_onE [where e = ee])
       
  7390           using \<open>0 < ee\<close> by auto
       
  7391         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"
       
  7392                  for  w z
       
  7393           using \<open>dd>0\<close> kk [of "(x,z)" "(w,z)"] by (force simp add: norm_minus_commute dist_norm)
       
  7394         show ?thesis
       
  7395           using ax unfolding lim_sequentially eventually_sequentially
       
  7396           apply (drule_tac x="min dd kk" in spec)
       
  7397           using \<open>dd > 0\<close> \<open>kk > 0\<close>
       
  7398           apply (fastforce simp: kk dist_norm)
       
  7399           done
       
  7400       qed
       
  7401       have tendsto_hx: "(\<lambda>n. contour_integral \<gamma> (d (a n))) \<longlonglongrightarrow> h x"
       
  7402         apply (simp add: contour_integral_unique [OF U, symmetric] x)
       
  7403         apply (rule contour_integral_uniform_limit [OF A1 A2 le_B])
       
  7404         apply (auto simp: \<open>valid_path \<gamma>\<close>)
       
  7405         done
       
  7406       then show "(h \<circ> a) \<longlonglongrightarrow> h x"
       
  7407         by (simp add: h_def x au o_def)
       
  7408     qed
       
  7409     show ?thesis
       
  7410     proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify)
       
  7411       fix z0
       
  7412       consider "z0 \<in> v" | "z0 \<in> u" using uv_Un by blast
       
  7413       then show "h field_differentiable at z0"
       
  7414       proof cases
       
  7415         assume "z0 \<in> v" then show ?thesis
       
  7416           using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \<open>valid_path \<gamma>\<close>
       
  7417           by (auto simp: field_differentiable_def v_def)
       
  7418       next
       
  7419         assume "z0 \<in> u" then
       
  7420         obtain e where "e>0" and e: "ball z0 e \<subseteq> u" by (blast intro: openE [OF u])
       
  7421         have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0"
       
  7422                 if abc_subset: "convex hull {a, b, c} \<subseteq> ball z0 e"  for a b c
       
  7423         proof -
       
  7424           have *: "\<And>x1 x2 z. z \<in> u \<Longrightarrow> closed_segment x1 x2 \<subseteq> u \<Longrightarrow> (\<lambda>w. d w z) contour_integrable_on linepath x1 x2"
       
  7425             using  hol_dw holomorphic_on_imp_continuous_on u
       
  7426             by (auto intro!: contour_integrable_holomorphic_simple)
       
  7427           have abc: "closed_segment a b \<subseteq> u"  "closed_segment b c \<subseteq> u"  "closed_segment c a \<subseteq> u"
       
  7428             using that e segments_subset_convex_hull by fastforce+
       
  7429           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"
       
  7430             apply (rule contour_integral_unique [OF Cauchy_theorem_triangle])
       
  7431             apply (rule holomorphic_on_subset [OF hol_dw])
       
  7432             using e abc_subset by auto
       
  7433           have "contour_integral \<gamma>
       
  7434                    (\<lambda>x. contour_integral (linepath a b) (\<lambda>z. d z x) +
       
  7435                         (contour_integral (linepath b c) (\<lambda>z. d z x) +
       
  7436                          contour_integral (linepath c a) (\<lambda>z. d z x)))  =  0"
       
  7437             apply (rule contour_integral_eq_0)
       
  7438             using abc pasz u
       
  7439             apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+
       
  7440             done
       
  7441           then show ?thesis
       
  7442             by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac)
       
  7443         qed
       
  7444         show ?thesis
       
  7445           using e \<open>e > 0\<close>
       
  7446           by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic
       
  7447                            Morera_triangle continuous_on_subset [OF conthu] *)
       
  7448       qed
       
  7449     qed
       
  7450   qed
       
  7451   ultimately have [simp]: "h z = 0" for z
       
  7452     by (meson Liouville_weak)
       
  7453   have "((\<lambda>w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z) \<gamma>"
       
  7454     by (rule has_contour_integral_winding_number [OF \<open>valid_path \<gamma>\<close> znot])
       
  7455   then have "((\<lambda>w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z * f z) \<gamma>"
       
  7456     by (metis mult.commute has_contour_integral_lmul)
       
  7457   then have 1: "((\<lambda>w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z * f z) \<gamma>"
       
  7458     by (simp add: divide_simps)
       
  7459   moreover have 2: "((\<lambda>w. (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>"
       
  7460     using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\<lambda>w. (f w - f z)/(w - z)"])
       
  7461   show ?thesis
       
  7462     using has_contour_integral_add [OF 1 2]  by (simp add: diff_divide_distrib)
       
  7463 qed
       
  7464 
       
  7465 
       
  7466 theorem Cauchy_integral_formula_global:
       
  7467     assumes s: "open s" and holf: "f holomorphic_on s"
       
  7468         and z: "z \<in> s" and vpg: "valid_path \<gamma>"
       
  7469         and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  7470         and zero: "\<And>w. w \<notin> s \<Longrightarrow> winding_number \<gamma> w = 0"
       
  7471       shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
       
  7472 proof -
       
  7473   have "path \<gamma>" using vpg by (blast intro: valid_path_imp_path)
       
  7474   have hols: "(\<lambda>w. f w / (w - z)) holomorphic_on s - {z}" "(\<lambda>w. 1 / (w - z)) holomorphic_on s - {z}"
       
  7475     by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+
       
  7476   then have cint_fw: "(\<lambda>w. f w / (w - z)) contour_integrable_on \<gamma>"
       
  7477     by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete s vpg pasz)
       
  7478   obtain d where "d>0"
       
  7479       and d: "\<And>g h. \<lbrakk>valid_path g; valid_path h; \<forall>t\<in>{0..1}. cmod (g t - \<gamma> t) < d \<and> cmod (h t - \<gamma> t) < d;
       
  7480                      pathstart h = pathstart g \<and> pathfinish h = pathfinish g\<rbrakk>
       
  7481                      \<Longrightarrow> path_image h \<subseteq> s - {z} \<and> (\<forall>f. f holomorphic_on s - {z} \<longrightarrow> contour_integral h f = contour_integral g f)"
       
  7482     using contour_integral_nearby_ends [OF _ \<open>path \<gamma>\<close> pasz] s by (simp add: open_Diff) metis
       
  7483   obtain p where polyp: "polynomial_function p"
       
  7484              and ps: "pathstart p = pathstart \<gamma>" and pf: "pathfinish p = pathfinish \<gamma>" and led: "\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < d"
       
  7485     using path_approx_polynomial_function [OF \<open>path \<gamma>\<close> \<open>d > 0\<close>] by blast
       
  7486   then have ploop: "pathfinish p = pathstart p" using loop by auto
       
  7487   have vpp: "valid_path p"  using polyp valid_path_polynomial_function by blast
       
  7488   have [simp]: "z \<notin> path_image \<gamma>" using pasz by blast
       
  7489   have paps: "path_image p \<subseteq> s - {z}" and cint_eq: "(\<And>f. f holomorphic_on s - {z} \<Longrightarrow> contour_integral p f = contour_integral \<gamma> f)"
       
  7490     using pf ps led d [OF vpg vpp] \<open>d > 0\<close> by auto
       
  7491   have wn_eq: "winding_number p z = winding_number \<gamma> z"
       
  7492     using vpp paps
       
  7493     by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols)
       
  7494   have "winding_number p w = winding_number \<gamma> w" if "w \<notin> s" for w
       
  7495   proof -
       
  7496     have hol: "(\<lambda>v. 1 / (v - w)) holomorphic_on s - {z}"
       
  7497       using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf])
       
  7498    have "w \<notin> path_image p" "w \<notin> path_image \<gamma>" using paps pasz that by auto
       
  7499    then show ?thesis
       
  7500     using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol])
       
  7501   qed
       
  7502   then have wn0: "\<And>w. w \<notin> s \<Longrightarrow> winding_number p w = 0"
       
  7503     by (simp add: zero)
       
  7504   show ?thesis
       
  7505     using Cauchy_integral_formula_global_weak [OF s holf z polyp paps ploop wn0] hols
       
  7506     by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq)
       
  7507 qed
       
  7508 
       
  7509 theorem Cauchy_theorem_global:
       
  7510     assumes s: "open s" and holf: "f holomorphic_on s"
       
  7511         and vpg: "valid_path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
       
  7512         and pas: "path_image \<gamma> \<subseteq> s"
       
  7513         and zero: "\<And>w. w \<notin> s \<Longrightarrow> winding_number \<gamma> w = 0"
       
  7514       shows "(f has_contour_integral 0) \<gamma>"
       
  7515 proof -
       
  7516   obtain z where "z \<in> s" and znot: "z \<notin> path_image \<gamma>"
       
  7517   proof -
       
  7518     have "compact (path_image \<gamma>)"
       
  7519       using compact_valid_path_image vpg by blast
       
  7520     then have "path_image \<gamma> \<noteq> s"
       
  7521       by (metis (no_types) compact_open path_image_nonempty s)
       
  7522     with pas show ?thesis by (blast intro: that)
       
  7523   qed
       
  7524   then have pasz: "path_image \<gamma> \<subseteq> s - {z}" using pas by blast
       
  7525   have hol: "(\<lambda>w. (w - z) * f w) holomorphic_on s"
       
  7526     by (rule holomorphic_intros holf)+
       
  7527   show ?thesis
       
  7528     using Cauchy_integral_formula_global [OF s hol \<open>z \<in> s\<close> vpg pasz loop zero]
       
  7529     by (auto simp: znot elim!: has_contour_integral_eq)
       
  7530 qed
       
  7531 
       
  7532 corollary Cauchy_theorem_global_outside:
       
  7533     assumes "open s" "f holomorphic_on s" "valid_path \<gamma>"  "pathfinish \<gamma> = pathstart \<gamma>" "path_image \<gamma> \<subseteq> s"
       
  7534             "\<And>w. w \<notin> s \<Longrightarrow> w \<in> outside(path_image \<gamma>)"
       
  7535       shows "(f has_contour_integral 0) \<gamma>"
       
  7536 by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path)
       
  7537 
       
  7538 
       
  7539 end