New complex analysis material
authorpaulson <lp15@cam.ac.uk>
Wed, 19 Mar 2014 14:54:45 +0000
changeset 56215 fcf90317383d
parent 56212 3253aaf73a01
child 56216 76ff0a15d202
New complex analysis material
src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy
src/HOL/Multivariate_Analysis/PolyRoots.thy
src/HOL/ROOT
src/HOL/Set_Interval.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy	Wed Mar 19 14:54:45 2014 +0000
@@ -0,0 +1,1445 @@
+(*  Author: John Harrison, Marco Maggesi, Graziano Gentili, Gianni Ciolli, Valentina Bruno
+    Ported from "hol_light/Multivariate/canal.ml" by L C Paulson (2014)
+*)
+
+header {* Complex Analysis Basics *}
+
+theory Complex_Analysis_Basics
+imports  "~~/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space"
+
+begin
+
+subsection {*Complex number lemmas *}
+
+lemma abs_sqrt_wlog:
+  fixes x::"'a::linordered_idom"
+  assumes "!!x::'a. x\<ge>0 \<Longrightarrow> P x (x\<^sup>2)" shows "P \<bar>x\<bar> (x\<^sup>2)"
+by (metis abs_ge_zero assms power2_abs)
+
+lemma complex_abs_le_norm: "abs(Re z) + abs(Im z) \<le> sqrt(2) * norm z"
+proof (cases z)
+  case (Complex x y)
+  show ?thesis
+    apply (rule power2_le_imp_le)
+    apply (auto simp: real_sqrt_mult [symmetric] Complex)
+    apply (rule abs_sqrt_wlog [where x=x])
+    apply (rule abs_sqrt_wlog [where x=y])
+    apply (simp add: power2_sum add_commute sum_squares_bound)
+    done
+qed
+
+lemma continuous_Re: "continuous_on UNIV Re"
+  by (metis (poly_guards_query) bounded_linear.continuous_on bounded_linear_Re 
+            continuous_on_cong continuous_on_id)
+
+lemma continuous_Im: "continuous_on UNIV Im"
+  by (metis (poly_guards_query) bounded_linear.continuous_on bounded_linear_Im 
+            continuous_on_cong continuous_on_id)
+
+lemma open_halfspace_Re_lt: "open {z. Re(z) < b}"
+proof -
+  have "{z. Re(z) < b} = Re -`{..<b}"
+    by blast
+  then show ?thesis
+    by (auto simp: continuous_Re continuous_imp_open_vimage [of UNIV])
+qed
+
+lemma open_halfspace_Re_gt: "open {z. Re(z) > b}"
+proof -
+  have "{z. Re(z) > b} = Re -`{b<..}"
+    by blast
+  then show ?thesis
+    by (auto simp: continuous_Re continuous_imp_open_vimage [of UNIV])
+qed
+
+lemma closed_halfspace_Re_ge: "closed {z. Re(z) \<ge> b}"
+proof -
+  have "{z. Re(z) \<ge> b} = - {z. Re(z) < b}"
+    by auto
+  then show ?thesis
+    by (simp add: closed_def open_halfspace_Re_lt)
+qed
+
+lemma closed_halfspace_Re_le: "closed {z. Re(z) \<le> b}"
+proof -
+  have "{z. Re(z) \<le> b} = - {z. Re(z) > b}"
+    by auto
+  then show ?thesis
+    by (simp add: closed_def open_halfspace_Re_gt)
+qed
+
+lemma closed_halfspace_Re_eq: "closed {z. Re(z) = b}"
+proof -
+  have "{z. Re(z) = b} = {z. Re(z) \<le> b} \<inter> {z. Re(z) \<ge> b}"
+    by auto
+  then show ?thesis
+    by (auto simp: closed_Int closed_halfspace_Re_le closed_halfspace_Re_ge)
+qed
+
+lemma open_halfspace_Im_lt: "open {z. Im(z) < b}"
+proof -
+  have "{z. Im(z) < b} = Im -`{..<b}"
+    by blast
+  then show ?thesis
+    by (auto simp: continuous_Im continuous_imp_open_vimage [of UNIV])
+qed
+
+lemma open_halfspace_Im_gt: "open {z. Im(z) > b}"
+proof -
+  have "{z. Im(z) > b} = Im -`{b<..}"
+    by blast
+  then show ?thesis
+    by (auto simp: continuous_Im continuous_imp_open_vimage [of UNIV])
+qed
+
+lemma closed_halfspace_Im_ge: "closed {z. Im(z) \<ge> b}"
+proof -
+  have "{z. Im(z) \<ge> b} = - {z. Im(z) < b}"
+    by auto
+  then show ?thesis
+    by (simp add: closed_def open_halfspace_Im_lt)
+qed
+
+lemma closed_halfspace_Im_le: "closed {z. Im(z) \<le> b}"
+proof -
+  have "{z. Im(z) \<le> b} = - {z. Im(z) > b}"
+    by auto
+  then show ?thesis
+    by (simp add: closed_def open_halfspace_Im_gt)
+qed
+
+lemma closed_halfspace_Im_eq: "closed {z. Im(z) = b}"
+proof -
+  have "{z. Im(z) = b} = {z. Im(z) \<le> b} \<inter> {z. Im(z) \<ge> b}"
+    by auto
+  then show ?thesis
+    by (auto simp: closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge)
+qed
+
+lemma complex_is_Real_iff: "z \<in> \<real> \<longleftrightarrow> Im z = 0"
+  by (metis Complex_in_Reals Im_complex_of_real Reals_cases complex_surj)
+
+lemma closed_complex_Reals: "closed (Reals :: complex set)"
+proof -
+  have "-(Reals :: complex set) = {z. Im(z) < 0} \<union> {z. 0 < Im(z)}"
+    by (auto simp: complex_is_Real_iff)
+  then show ?thesis
+    by (metis closed_def open_Un open_halfspace_Im_gt open_halfspace_Im_lt)
+qed
+
+
+lemma linear_times:
+  fixes c::"'a::{real_algebra,real_vector}" shows "linear (\<lambda>x. c * x)"
+  by (auto simp: linearI distrib_left)
+
+lemma bilinear_times:
+  fixes c::"'a::{real_algebra,real_vector}" shows "bilinear (\<lambda>x y::'a. x*y)"
+  unfolding bilinear_def
+  by (auto simp: distrib_left distrib_right intro!: linearI)
+
+lemma linear_cnj: "linear cnj"
+  by (auto simp: linearI cnj_def)
+
+lemma tendsto_mult_left:
+  fixes c::"'a::real_normed_algebra" 
+  shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) ---> c * l) F"
+by (rule tendsto_mult [OF tendsto_const])
+
+lemma tendsto_mult_right:
+  fixes c::"'a::real_normed_algebra" 
+  shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) ---> l * c) F"
+by (rule tendsto_mult [OF _ tendsto_const])
+
+lemma tendsto_Re_upper:
+  assumes "~ (trivial_limit F)" 
+          "(f ---> l) F" 
+          "eventually (\<lambda>x. Re(f x) \<le> b) F"
+    shows  "Re(l) \<le> b"
+  by (metis assms tendsto_le [OF _ tendsto_const]  tendsto_Re)
+
+lemma tendsto_Re_lower:
+  assumes "~ (trivial_limit F)" 
+          "(f ---> l) F" 
+          "eventually (\<lambda>x. b \<le> Re(f x)) F"
+    shows  "b \<le> Re(l)"
+  by (metis assms tendsto_le [OF _ _ tendsto_const]  tendsto_Re)
+
+lemma tendsto_Im_upper:
+  assumes "~ (trivial_limit F)" 
+          "(f ---> l) F" 
+          "eventually (\<lambda>x. Im(f x) \<le> b) F"
+    shows  "Im(l) \<le> b"
+  by (metis assms tendsto_le [OF _ tendsto_const]  tendsto_Im)
+
+lemma tendsto_Im_lower:
+  assumes "~ (trivial_limit F)" 
+          "(f ---> l) F" 
+          "eventually (\<lambda>x. b \<le> Im(f x)) F"
+    shows  "b \<le> Im(l)"
+  by (metis assms tendsto_le [OF _ _ tendsto_const]  tendsto_Im)
+
+subsection{*General lemmas*}
+
+lemma continuous_mult_left:
+  fixes c::"'a::real_normed_algebra" 
+  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. c * f x)"
+by (rule continuous_mult [OF continuous_const])
+
+lemma continuous_mult_right:
+  fixes c::"'a::real_normed_algebra" 
+  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x * c)"
+by (rule continuous_mult [OF _ continuous_const])
+
+lemma continuous_on_mult_left:
+  fixes c::"'a::real_normed_algebra" 
+  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. c * f x)"
+by (rule continuous_on_mult [OF continuous_on_const])
+
+lemma continuous_on_mult_right:
+  fixes c::"'a::real_normed_algebra" 
+  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x * c)"
+by (rule continuous_on_mult [OF _ continuous_on_const])
+
+lemma uniformly_continuous_on_cmul_right [continuous_on_intros]:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"
+  assumes "uniformly_continuous_on s f"
+  shows "uniformly_continuous_on s (\<lambda>x. f x * c)"
+proof (cases "c=0")
+  case True then show ?thesis
+    by (simp add: uniformly_continuous_on_const)
+next
+  case False show ?thesis
+    apply (rule bounded_linear.uniformly_continuous_on)
+    apply (metis bounded_linear_ident)
+    using assms
+    apply (auto simp: uniformly_continuous_on_def dist_norm)
+    apply (drule_tac x = "e / norm c" in spec, auto)
+    apply (metis divide_pos_pos zero_less_norm_iff False)
+    apply (rule_tac x=d in exI, auto)
+    apply (drule_tac x = x in bspec, assumption)
+    apply (drule_tac x = "x'" in bspec)
+    apply (auto simp: False less_divide_eq)
+    by (metis dual_order.strict_trans2 left_diff_distrib norm_mult_ineq)
+qed
+
+lemma uniformly_continuous_on_cmul_left[continuous_on_intros]:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"
+  assumes "uniformly_continuous_on s f"
+    shows "uniformly_continuous_on s (\<lambda>x. c * f x)"
+by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right)
+
+lemma continuous_within_norm_id [continuous_intros]: "continuous (at x within S) norm"
+  by (rule continuous_norm [OF continuous_ident])
+
+lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm"
+  by (metis continuous_on_eq continuous_on_id continuous_on_norm)
+
+
+subsection{*DERIV stuff*}
+
+(*move some premises to a sensible order. Use more \<And> symbols.*)
+
+lemma DERIV_continuous_on: "\<lbrakk>\<And>x. x \<in> s \<Longrightarrow> DERIV f x :> D\<rbrakk> \<Longrightarrow> continuous_on s f"
+  by (metis DERIV_continuous continuous_at_imp_continuous_on)
+
+lemma DERIV_subset: 
+  "(f has_field_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s 
+   \<Longrightarrow> (f has_field_derivative f') (at x within t)"
+  by (simp add: has_field_derivative_def has_derivative_within_subset)
+
+lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = op * 0"
+  by auto
+
+lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = op * 1"
+  by auto
+
+lemma has_derivative_zero_constant:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "convex s"
+      and d0: "\<And>x. x\<in>s \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within s)"
+    shows "\<exists>c. \<forall>x\<in>s. f x = c"
+proof (cases "s={}")
+  case False
+  then obtain x where "x \<in> s"
+    by auto
+  have d0': "\<forall>x\<in>s. (f has_derivative (\<lambda>h. 0)) (at x within s)"
+    by (metis d0)
+  have "\<And>y. y \<in> s \<Longrightarrow> f x = f y"
+  proof -
+    case goal1
+    then show ?case
+      using differentiable_bound[OF assms(1) d0', of 0 x y] and `x \<in> s`
+      unfolding onorm_const
+      by auto
+  qed
+  then show ?thesis 
+    by metis
+next
+  case True
+  then show ?thesis by auto
+qed
+
+lemma has_derivative_zero_unique:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "convex s"
+      and "\<And>x. x\<in>s \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within s)"
+      and "a \<in> s"
+      and "x \<in> s"
+    shows "f x = f a"
+  using assms
+  by (metis has_derivative_zero_constant)
+
+lemma has_derivative_zero_connected_constant:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  assumes "connected s"
+      and "open s"
+      and "finite k"
+      and "continuous_on s f"
+      and "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)"
+    obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c"
+proof (cases "s = {}")
+  case True
+  then show ?thesis
+by (metis empty_iff that)
+next
+  case False
+  then obtain c where "c \<in> s"
+    by (metis equals0I)
+  then show ?thesis
+    by (metis has_derivative_zero_unique_strong_connected assms that)
+qed
+
+lemma DERIV_zero_connected_constant:
+  fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
+  assumes "connected s"
+      and "open s"
+      and "finite k"
+      and "continuous_on s f"
+      and "\<forall>x\<in>(s - k). DERIV f x :> 0"
+    obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c"
+using has_derivative_zero_connected_constant [OF assms(1-4)] assms
+by (metis DERIV_const Derivative.has_derivative_const Diff_iff at_within_open 
+          frechet_derivative_at has_field_derivative_def)
+
+lemma DERIV_zero_constant:
+  fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
+  shows    "\<lbrakk>convex s;
+             \<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)\<rbrakk> 
+             \<Longrightarrow> \<exists>c. \<forall>x \<in> s. f(x) = c"
+  unfolding has_field_derivative_def
+  by (auto simp: lambda_zero intro: has_derivative_zero_constant)
+
+lemma DERIV_zero_unique:
+  fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
+  assumes "convex s"
+      and d0: "\<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)"
+      and "a \<in> s"
+      and "x \<in> s"
+    shows "f x = f a"
+apply (rule has_derivative_zero_unique [where f=f, OF assms(1) _ assms(3,4)])
+by (metis d0 has_field_derivative_imp_has_derivative lambda_zero)
+
+lemma DERIV_zero_connected_unique:
+  fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
+  assumes "connected s"
+      and "open s"
+      and d0: "\<And>x. x\<in>s \<Longrightarrow> DERIV f x :> 0"
+      and "a \<in> s"
+      and "x \<in> s"
+    shows "f x = f a" 
+    apply (rule Integration.has_derivative_zero_unique_strong_connected [of s "{}" f])
+    using assms
+    apply auto
+    apply (metis DERIV_continuous_on)
+  by (metis at_within_open has_field_derivative_def lambda_zero)
+
+lemma DERIV_transform_within:
+  assumes "(f has_field_derivative f') (at a within s)"
+      and "0 < d" "a \<in> s"
+      and "\<And>x. x\<in>s \<Longrightarrow> dist x a < d \<Longrightarrow> f x = g x"
+    shows "(g has_field_derivative f') (at a within s)"
+  using assms unfolding has_field_derivative_def
+  by (blast intro: Derivative.has_derivative_transform_within)
+
+lemma DERIV_transform_within_open:
+  assumes "DERIV f a :> f'"
+      and "open s" "a \<in> s"
+      and "\<And>x. x\<in>s \<Longrightarrow> f x = g x"
+    shows "DERIV g a :> f'"
+  using assms unfolding has_field_derivative_def
+by (metis has_derivative_transform_within_open)
+
+lemma DERIV_transform_at:
+  assumes "DERIV f a :> f'"
+      and "0 < d"
+      and "\<And>x. dist x a < d \<Longrightarrow> f x = g x"
+    shows "DERIV g a :> f'"
+  by (blast intro: assms DERIV_transform_within)
+
+
+subsection{*Holomorphic functions*}
+
+lemma has_derivative_ident[has_derivative_intros, simp]: 
+     "FDERIV complex_of_real x :> complex_of_real"
+  by (simp add: has_derivative_def tendsto_const bounded_linear_of_real)
+
+lemma has_real_derivative:
+  fixes f :: "real\<Rightarrow>real" 
+  assumes "(f has_derivative f') F"
+    obtains c where "(f has_derivative (\<lambda>x. x * c)) F"
+proof -
+  obtain c where "f' = (\<lambda>x. x * c)"
+    by (metis assms derivative_linear real_bounded_linear)
+  then show ?thesis
+    by (metis assms that)
+qed
+
+lemma has_real_derivative_iff:
+  fixes f :: "real\<Rightarrow>real" 
+  shows "(\<exists>f'. (f has_derivative (\<lambda>x. x * f')) F) = (\<exists>D. (f has_derivative D) F)"
+by (auto elim: has_real_derivative)
+
+definition complex_differentiable :: "[complex \<Rightarrow> complex, complex filter] \<Rightarrow> bool"
+           (infixr "(complex'_differentiable)" 50)  
+  where "f complex_differentiable F \<equiv> \<exists>f'. (f has_field_derivative f') F"
+
+definition DD :: "['a \<Rightarrow> 'a::real_normed_field, 'a] \<Rightarrow> 'a" --{*for real, complex?*}
+  where "DD f x \<equiv> THE f'. (f has_derivative (\<lambda>x. x * f')) (at x)"
+
+definition holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool"
+           (infixl "(holomorphic'_on)" 50)
+  where "f holomorphic_on s \<equiv> \<forall>x \<in> s. \<exists>f'. (f has_field_derivative f') (at x within s)"
+  
+lemma holomorphic_on_empty: "f holomorphic_on {}"
+  by (simp add: holomorphic_on_def)
+
+lemma holomorphic_on_differentiable:
+     "f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. f complex_differentiable (at x within s))"
+unfolding holomorphic_on_def complex_differentiable_def has_field_derivative_def
+by (metis mult_commute_abs)
+
+lemma holomorphic_on_open:
+    "open s \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>f'. DERIV f x :> f')"
+  by (auto simp: holomorphic_on_def has_field_derivative_def at_within_open [of _ s])
+
+lemma complex_differentiable_imp_continuous_at: 
+    "f complex_differentiable (at x) \<Longrightarrow> continuous (at x) f"
+  by (metis DERIV_continuous complex_differentiable_def)
+
+lemma holomorphic_on_imp_continuous_on: 
+    "f holomorphic_on s \<Longrightarrow> continuous_on s f"
+  by (metis DERIV_continuous continuous_on_eq_continuous_within holomorphic_on_def) 
+
+lemma has_derivative_within_open:
+  "a \<in> s \<Longrightarrow> open s \<Longrightarrow> (f has_field_derivative f') (at a within s) \<longleftrightarrow> DERIV f a :> f'"
+  by (simp add: has_field_derivative_def) (metis has_derivative_within_open)
+
+lemma holomorphic_on_subset:
+    "f holomorphic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f holomorphic_on t"
+  unfolding holomorphic_on_def
+  by (metis DERIV_subset subsetD)
+
+lemma complex_differentiable_within_subset:
+    "\<lbrakk>f complex_differentiable (at x within s); t \<subseteq> s\<rbrakk>
+     \<Longrightarrow> f complex_differentiable (at x within t)"
+  unfolding complex_differentiable_def
+  by (metis DERIV_subset)
+
+lemma complex_differentiable_at_within:
+    "\<lbrakk>f complex_differentiable (at x)\<rbrakk>
+     \<Longrightarrow> f complex_differentiable (at x within s)"
+  unfolding complex_differentiable_def
+  by (metis DERIV_subset top_greatest)
+
+
+lemma has_derivative_mult_right:
+  fixes c:: "'a :: real_normed_algebra"
+  shows "((op * c) has_derivative (op * c)) F"
+by (rule has_derivative_mult_right [OF has_derivative_id])
+
+lemma complex_differentiable_linear:
+     "(op * c) complex_differentiable F"
+proof -
+  have "\<And>u::complex. (\<lambda>x. x * u) = op * u"
+    by (rule ext) (simp add: mult_ac)
+  then show ?thesis
+    unfolding complex_differentiable_def has_field_derivative_def
+    by (force intro: has_derivative_mult_right)
+qed
+
+lemma complex_differentiable_const:
+  "(\<lambda>z. c) complex_differentiable F"
+  unfolding complex_differentiable_def has_field_derivative_def
+  apply (rule exI [where x=0])
+  by (metis Derivative.has_derivative_const lambda_zero) 
+
+lemma complex_differentiable_id:
+  "(\<lambda>z. z) complex_differentiable F"
+  unfolding complex_differentiable_def has_field_derivative_def
+  apply (rule exI [where x=1])
+  apply (simp add: lambda_one [symmetric])
+  done
+
+(*DERIV_minus*)
+lemma field_differentiable_minus:
+  assumes "(f has_field_derivative f') F" 
+    shows "((\<lambda>z. - (f z)) has_field_derivative -f') F"
+  apply (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus])
+  using assms 
+  by (auto simp: has_field_derivative_def field_simps mult_commute_abs)
+
+(*DERIV_add*)
+lemma field_differentiable_add:
+  assumes "(f has_field_derivative f') F" "(g has_field_derivative g') F"
+    shows "((\<lambda>z. f z + g z) has_field_derivative f' + g') F"
+  apply (rule has_derivative_imp_has_field_derivative[OF has_derivative_add])
+  using assms 
+  by (auto simp: has_field_derivative_def field_simps mult_commute_abs)
+
+(*DERIV_diff*)
+lemma field_differentiable_diff:
+  assumes "(f has_field_derivative f') F" "(g has_field_derivative g') F"
+    shows "((\<lambda>z. f z - g z) has_field_derivative f' - g') F"
+by (simp only: assms diff_conv_add_uminus field_differentiable_add field_differentiable_minus)
+
+lemma complex_differentiable_minus:
+    "f complex_differentiable F \<Longrightarrow> (\<lambda>z. -(f z)) complex_differentiable F"
+  using assms unfolding complex_differentiable_def
+  by (metis field_differentiable_minus)
+
+lemma complex_differentiable_add:
+  assumes "f complex_differentiable F" "g complex_differentiable F"
+    shows "(\<lambda>z. f z + g z) complex_differentiable F"
+  using assms unfolding complex_differentiable_def
+  by (metis field_differentiable_add)
+
+lemma complex_differentiable_diff:
+  assumes "f complex_differentiable F" "g complex_differentiable F"
+    shows "(\<lambda>z. f z - g z) complex_differentiable F"
+  using assms unfolding complex_differentiable_def
+  by (metis field_differentiable_diff)
+
+lemma complex_differentiable_inverse:
+  assumes "f complex_differentiable (at a within s)" "f a \<noteq> 0"
+  shows "(\<lambda>z. inverse (f z)) complex_differentiable (at a within s)"
+  using assms unfolding complex_differentiable_def
+  by (metis DERIV_inverse_fun)
+
+lemma complex_differentiable_mult:
+  assumes "f complex_differentiable (at a within s)" 
+          "g complex_differentiable (at a within s)"
+    shows "(\<lambda>z. f z * g z) complex_differentiable (at a within s)"
+  using assms unfolding complex_differentiable_def
+  by (metis DERIV_mult [of f _ a s g])
+  
+lemma complex_differentiable_divide:
+  assumes "f complex_differentiable (at a within s)" 
+          "g complex_differentiable (at a within s)"
+          "g a \<noteq> 0"
+    shows "(\<lambda>z. f z / g z) complex_differentiable (at a within s)"
+  using assms unfolding complex_differentiable_def
+  by (metis DERIV_divide [of f _ a s g])
+
+lemma complex_differentiable_power:
+  assumes "f complex_differentiable (at a within s)" 
+    shows "(\<lambda>z. f z ^ n) complex_differentiable (at a within s)"
+  using assms unfolding complex_differentiable_def
+  by (metis DERIV_power)
+
+lemma complex_differentiable_transform_within:
+  "0 < d \<Longrightarrow>
+        x \<in> s \<Longrightarrow>
+        (\<And>x'. x' \<in> s \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x') \<Longrightarrow>
+        f complex_differentiable (at x within s)
+        \<Longrightarrow> g complex_differentiable (at x within s)"
+  unfolding complex_differentiable_def has_field_derivative_def
+  by (blast intro: has_derivative_transform_within)
+
+lemma complex_differentiable_compose_within:
+  assumes "f complex_differentiable (at a within s)" 
+          "g complex_differentiable (at (f a) within f`s)"
+    shows "(g o f) complex_differentiable (at a within s)"
+  using assms unfolding complex_differentiable_def
+  by (metis DERIV_image_chain)
+
+lemma complex_differentiable_within_open:
+     "\<lbrakk>a \<in> s; open s\<rbrakk> \<Longrightarrow> f complex_differentiable at a within s \<longleftrightarrow> 
+                          f complex_differentiable at a"
+  unfolding complex_differentiable_def
+  by (metis at_within_open)
+
+lemma holomorphic_transform:
+     "\<lbrakk>f holomorphic_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g holomorphic_on s"
+  apply (auto simp: holomorphic_on_def has_field_derivative_def)
+  by (metis complex_differentiable_def complex_differentiable_transform_within has_field_derivative_def linordered_field_no_ub)
+
+lemma holomorphic_eq:
+     "(\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> g holomorphic_on s"
+  by (metis holomorphic_transform)
+
+subsection{*Holomorphic*}
+
+lemma holomorphic_on_linear:
+     "(op * c) holomorphic_on s"
+  unfolding holomorphic_on_def  by (metis DERIV_cmult_Id)
+
+lemma holomorphic_on_const:
+     "(\<lambda>z. c) holomorphic_on s"
+  unfolding holomorphic_on_def  
+  by (metis DERIV_const)
+
+lemma holomorphic_on_id:
+     "id holomorphic_on s"
+  unfolding holomorphic_on_def id_def  
+  by (metis DERIV_ident)
+
+lemma holomorphic_on_compose:
+  "f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s)
+           \<Longrightarrow> (g o f) holomorphic_on s"
+  unfolding holomorphic_on_def
+  by (metis DERIV_image_chain imageI)
+
+lemma holomorphic_on_compose_gen:
+  "\<lbrakk>f holomorphic_on s; g holomorphic_on t; f ` s \<subseteq> t\<rbrakk> \<Longrightarrow> (g o f) holomorphic_on s"
+  unfolding holomorphic_on_def
+  by (metis DERIV_image_chain DERIV_subset image_subset_iff)
+
+lemma holomorphic_on_minus:
+  "f holomorphic_on s \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on s"
+  unfolding holomorphic_on_def
+by (metis DERIV_minus)
+
+lemma holomorphic_on_add:
+  "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) holomorphic_on s"
+  unfolding holomorphic_on_def
+  by (metis DERIV_add)
+
+lemma holomorphic_on_diff:
+  "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) holomorphic_on s"
+  unfolding holomorphic_on_def
+  by (metis DERIV_diff)
+
+lemma holomorphic_on_mult:
+  "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) holomorphic_on s"
+  unfolding holomorphic_on_def
+  by auto (metis DERIV_mult)
+
+lemma holomorphic_on_inverse:
+  "\<lbrakk>f holomorphic_on s; \<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. inverse (f z)) holomorphic_on s"
+  unfolding holomorphic_on_def
+  by (metis DERIV_inverse')
+
+lemma holomorphic_on_divide:
+  "\<lbrakk>f holomorphic_on s; g holomorphic_on s; \<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. f z / g z) holomorphic_on s"
+  unfolding holomorphic_on_def
+  by (metis (full_types) DERIV_divide)
+
+lemma holomorphic_on_power:
+  "f holomorphic_on s \<Longrightarrow> (\<lambda>z. (f z)^n) holomorphic_on s"
+  unfolding holomorphic_on_def
+  by (metis DERIV_power)
+
+lemma holomorphic_on_setsum:
+  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s)
+           \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) holomorphic_on s"
+  unfolding holomorphic_on_def
+  apply (induct I rule: finite_induct) 
+  apply (force intro: DERIV_const DERIV_add)+
+  done
+
+lemma DERIV_imp_DD: "DERIV f x :> f' \<Longrightarrow> DD f x = f'"
+    apply (simp add: DD_def has_field_derivative_def mult_commute_abs)
+    apply (rule the_equality, assumption)
+    apply (metis DERIV_unique has_field_derivative_def)
+    done
+
+lemma DD_iff_derivative_differentiable:
+  fixes f :: "real\<Rightarrow>real" 
+  shows   "DERIV f x :> DD f x \<longleftrightarrow> f differentiable at x"
+unfolding DD_def differentiable_def 
+by (metis (full_types) DD_def DERIV_imp_DD has_field_derivative_def has_real_derivative_iff 
+          mult_commute_abs)
+
+lemma DD_iff_derivative_complex_differentiable:
+  fixes f :: "complex\<Rightarrow>complex" 
+  shows "DERIV f x :> DD f x \<longleftrightarrow> f complex_differentiable at x"
+unfolding DD_def complex_differentiable_def
+by (metis DD_def DERIV_imp_DD)
+
+lemma complex_differentiable_compose:
+  "f complex_differentiable at z \<Longrightarrow> g complex_differentiable at (f z)
+          \<Longrightarrow> (g o f) complex_differentiable at z"
+by (metis complex_differentiable_at_within complex_differentiable_compose_within)
+
+lemma complex_derivative_chain:
+  fixes z::complex
+  shows
+  "f complex_differentiable at z \<Longrightarrow> g complex_differentiable at (f z)
+          \<Longrightarrow> DD (g o f) z = DD g (f z) * DD f z"
+by (metis DD_iff_derivative_complex_differentiable DERIV_chain DERIV_imp_DD)
+
+lemma comp_derivative_chain:
+  fixes z::real
+  shows "\<lbrakk>f differentiable at z; g differentiable at (f z)\<rbrakk> 
+         \<Longrightarrow> DD (g o f) z = DD g (f z) * DD f z"
+by (metis DD_iff_derivative_differentiable DERIV_chain DERIV_imp_DD)
+
+lemma complex_derivative_linear: "DD (\<lambda>w. c * w) = (\<lambda>z. c)"
+by (metis DERIV_imp_DD DERIV_cmult_Id)
+
+lemma complex_derivative_ident: "DD (\<lambda>w. w) = (\<lambda>z. 1)"
+by (metis DERIV_imp_DD DERIV_ident)
+
+lemma complex_derivative_const: "DD (\<lambda>w. c) = (\<lambda>z. 0)"
+by (metis DERIV_imp_DD DERIV_const)
+
+lemma complex_derivative_add:
+  "\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk>  
+   \<Longrightarrow> DD (\<lambda>w. f w + g w) z = DD f z + DD g z"
+  unfolding complex_differentiable_def
+  by (rule DERIV_imp_DD) (metis (poly_guards_query) DERIV_add DERIV_imp_DD)  
+
+lemma complex_derivative_diff:
+  "\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk>  
+   \<Longrightarrow> DD (\<lambda>w. f w - g w) z = DD f z - DD g z"
+  unfolding complex_differentiable_def
+  by (rule DERIV_imp_DD) (metis (poly_guards_query) DERIV_diff DERIV_imp_DD)
+
+lemma complex_derivative_mult:
+  "\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk>  
+   \<Longrightarrow> DD (\<lambda>w. f w * g w) z = f z * DD g z + DD f z * g z"
+  unfolding complex_differentiable_def
+  by (rule DERIV_imp_DD) (metis DERIV_imp_DD DERIV_mult')
+
+lemma complex_derivative_cmult:
+  "f complex_differentiable at z \<Longrightarrow> DD (\<lambda>w. c * f w) z = c * DD f z"
+  unfolding complex_differentiable_def
+  by (metis DERIV_cmult DERIV_imp_DD)
+
+lemma complex_derivative_cmult_right:
+  "f complex_differentiable at z \<Longrightarrow> DD (\<lambda>w. f w * c) z = DD f z * c"
+  unfolding complex_differentiable_def
+  by (metis DERIV_cmult_right DERIV_imp_DD)
+
+lemma complex_derivative_transform_within_open:
+  "\<lbrakk>f holomorphic_on s; g holomorphic_on s; open s; z \<in> s; \<And>w. w \<in> s \<Longrightarrow> f w = g w\<rbrakk> 
+   \<Longrightarrow> DD f z = DD g z"
+  unfolding holomorphic_on_def
+  by (rule DERIV_imp_DD) (metis has_derivative_within_open DERIV_imp_DD DERIV_transform_within_open)
+
+lemma complex_derivative_compose_linear:
+    "f complex_differentiable at (c * z) \<Longrightarrow> DD (\<lambda>w. f (c * w)) z = c * DD f (c * z)"
+apply (rule DERIV_imp_DD)
+apply (simp add: DD_iff_derivative_complex_differentiable [symmetric])
+apply (metis DERIV_chain' DERIV_cmult_Id comm_semiring_1_class.normalizing_semiring_rules(7))  
+done
+
+subsection{*Caratheodory characterization.*}
+
+(*REPLACE the original version. BUT IN WHICH FILE??*)
+thm Deriv.CARAT_DERIV
+
+lemma complex_differentiable_caratheodory_at:
+  "f complex_differentiable (at z) \<longleftrightarrow>
+         (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z) g)"
+  using CARAT_DERIV [of f]
+  by (simp add: complex_differentiable_def has_field_derivative_def)
+
+lemma complex_differentiable_caratheodory_within:
+  "f complex_differentiable (at z within s) \<longleftrightarrow>
+         (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z within s) g)"
+  using DERIV_caratheodory_within [of f]
+  by (simp add: complex_differentiable_def has_field_derivative_def)
+
+subsection{*analyticity on a set*}
+
+definition analytic_on (infixl "(analytic'_on)" 50)  
+  where
+   "f analytic_on s \<equiv> \<forall>x \<in> s. \<exists>e. 0 < e \<and> f holomorphic_on (ball x e)"
+
+lemma analytic_imp_holomorphic:
+     "f analytic_on s \<Longrightarrow> f holomorphic_on s"
+  unfolding analytic_on_def holomorphic_on_def
+  apply (simp add: has_derivative_within_open [OF _ open_ball])
+  apply (metis DERIV_subset dist_self mem_ball top_greatest)
+  done
+
+lemma analytic_on_open:
+     "open s \<Longrightarrow> f analytic_on s \<longleftrightarrow> f holomorphic_on s"
+apply (auto simp: analytic_imp_holomorphic)
+apply (auto simp: analytic_on_def holomorphic_on_def)
+by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball)
+
+lemma analytic_on_imp_differentiable_at:
+  "f analytic_on s \<Longrightarrow> x \<in> s \<Longrightarrow> f complex_differentiable (at x)"
+ apply (auto simp: analytic_on_def holomorphic_on_differentiable)
+by (metis Topology_Euclidean_Space.open_ball centre_in_ball complex_differentiable_within_open)
+
+lemma analytic_on_subset:
+     "f analytic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f analytic_on t"
+  by (auto simp: analytic_on_def)
+
+lemma analytic_on_Un:
+    "f analytic_on (s \<union> t) \<longleftrightarrow> f analytic_on s \<and> f analytic_on t"
+  by (auto simp: analytic_on_def)
+
+lemma analytic_on_Union:
+  "f analytic_on (\<Union> s) \<longleftrightarrow> (\<forall>t \<in> s. f analytic_on t)"
+  by (auto simp: analytic_on_def)
+  
+lemma analytic_on_holomorphic:
+  "f analytic_on s \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f holomorphic_on t)"
+  (is "?lhs = ?rhs")
+proof -
+  have "?lhs \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t)"
+  proof safe
+    assume "f analytic_on s"
+    then show "\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t"
+      apply (simp add: analytic_on_def)
+      apply (rule exI [where x="\<Union>{u. open u \<and> f analytic_on u}"], auto)
+      apply (metis Topology_Euclidean_Space.open_ball analytic_on_open centre_in_ball)
+      by (metis analytic_on_def)
+  next
+    fix t
+    assume "open t" "s \<subseteq> t" "f analytic_on t" 
+    then show "f analytic_on s"
+        by (metis analytic_on_subset)
+  qed
+  also have "... \<longleftrightarrow> ?rhs"
+    by (auto simp: analytic_on_open)
+  finally show ?thesis .
+qed
+
+lemma analytic_on_linear: "(op * c) analytic_on s"
+  apply (simp add: analytic_on_holomorphic holomorphic_on_linear)
+  by (metis open_UNIV top_greatest)
+
+lemma analytic_on_const: "(\<lambda>z. c) analytic_on s"
+  unfolding analytic_on_def
+  by (metis holomorphic_on_const zero_less_one)
+
+lemma analytic_on_id: "id analytic_on s"
+  unfolding analytic_on_def
+  apply (simp add: holomorphic_on_id)
+  by (metis gt_ex)
+
+lemma analytic_on_compose:
+  assumes f: "f analytic_on s"
+      and g: "g analytic_on (f ` s)"
+    shows "(g o f) analytic_on s"
+unfolding analytic_on_def
+proof (intro ballI)
+  fix x
+  assume x: "x \<in> s"
+  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f
+    by (metis analytic_on_def)
+  obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g
+    by (metis analytic_on_def g image_eqI x) 
+  have "isCont f x"
+    by (metis analytic_on_imp_differentiable_at complex_differentiable_imp_continuous_at f x)
+  with e' obtain d where d: "0 < d" and fd: "f ` ball x d \<subseteq> ball (f x) e'"
+     by (auto simp: continuous_at_ball)
+  have "g \<circ> f holomorphic_on ball x (min d e)" 
+    apply (rule holomorphic_on_compose)
+    apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+    by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball)
+  then show "\<exists>e>0. g \<circ> f holomorphic_on ball x e"
+    by (metis d e min_less_iff_conj) 
+qed
+
+lemma analytic_on_compose_gen:
+  "f analytic_on s \<Longrightarrow> g analytic_on t \<Longrightarrow> (\<And>z. z \<in> s \<Longrightarrow> f z \<in> t)
+             \<Longrightarrow> g o f analytic_on s"
+by (metis analytic_on_compose analytic_on_subset image_subset_iff)
+
+lemma analytic_on_neg:
+  "f analytic_on s \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on s"
+by (metis analytic_on_holomorphic holomorphic_on_minus)
+
+lemma analytic_on_add:
+  assumes f: "f analytic_on s"
+      and g: "g analytic_on s"
+    shows "(\<lambda>z. f z + g z) analytic_on s"
+unfolding analytic_on_def
+proof (intro ballI)
+  fix z
+  assume z: "z \<in> s"
+  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
+    by (metis analytic_on_def)
+  obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
+    by (metis analytic_on_def g z) 
+  have "(\<lambda>z. f z + g z) holomorphic_on ball z (min e e')" 
+    apply (rule holomorphic_on_add) 
+    apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+    by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+  then show "\<exists>e>0. (\<lambda>z. f z + g z) holomorphic_on ball z e"
+    by (metis e e' min_less_iff_conj)
+qed
+
+lemma analytic_on_diff:
+  assumes f: "f analytic_on s"
+      and g: "g analytic_on s"
+    shows "(\<lambda>z. f z - g z) analytic_on s"
+unfolding analytic_on_def
+proof (intro ballI)
+  fix z
+  assume z: "z \<in> s"
+  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
+    by (metis analytic_on_def)
+  obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
+    by (metis analytic_on_def g z) 
+  have "(\<lambda>z. f z - g z) holomorphic_on ball z (min e e')" 
+    apply (rule holomorphic_on_diff) 
+    apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+    by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+  then show "\<exists>e>0. (\<lambda>z. f z - g z) holomorphic_on ball z e"
+    by (metis e e' min_less_iff_conj)
+qed
+
+lemma analytic_on_mult:
+  assumes f: "f analytic_on s"
+      and g: "g analytic_on s"
+    shows "(\<lambda>z. f z * g z) analytic_on s"
+unfolding analytic_on_def
+proof (intro ballI)
+  fix z
+  assume z: "z \<in> s"
+  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
+    by (metis analytic_on_def)
+  obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
+    by (metis analytic_on_def g z) 
+  have "(\<lambda>z. f z * g z) holomorphic_on ball z (min e e')" 
+    apply (rule holomorphic_on_mult) 
+    apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+    by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
+  then show "\<exists>e>0. (\<lambda>z. f z * g z) holomorphic_on ball z e"
+    by (metis e e' min_less_iff_conj)
+qed
+
+lemma analytic_on_inverse:
+  assumes f: "f analytic_on s"
+      and nz: "(\<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0)"
+    shows "(\<lambda>z. inverse (f z)) analytic_on s"
+unfolding analytic_on_def
+proof (intro ballI)
+  fix z
+  assume z: "z \<in> s"
+  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
+    by (metis analytic_on_def)
+  have "continuous_on (ball z e) f"
+    by (metis fh holomorphic_on_imp_continuous_on)
+  then obtain e' where e': "0 < e'" and nz': "\<And>y. dist z y < e' \<Longrightarrow> f y \<noteq> 0" 
+    by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz)  
+  have "(\<lambda>z. inverse (f z)) holomorphic_on ball z (min e e')" 
+    apply (rule holomorphic_on_inverse)
+    apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball)
+    by (metis nz' mem_ball min_less_iff_conj) 
+  then show "\<exists>e>0. (\<lambda>z. inverse (f z)) holomorphic_on ball z e"
+    by (metis e e' min_less_iff_conj)
+qed
+
+
+lemma analytic_on_divide:
+  assumes f: "f analytic_on s"
+      and g: "g analytic_on s"
+      and nz: "(\<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0)"
+    shows "(\<lambda>z. f z / g z) analytic_on s"
+unfolding divide_inverse
+by (metis analytic_on_inverse analytic_on_mult f g nz)
+
+lemma analytic_on_power:
+  "f analytic_on s \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on s"
+by (induct n) (auto simp: analytic_on_const analytic_on_mult)
+
+lemma analytic_on_setsum:
+  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s)
+           \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) analytic_on s"
+  by (induct I rule: finite_induct) (auto simp: analytic_on_const analytic_on_add)
+
+subsection{*analyticity at a point.*}
+
+lemma analytic_at_ball:
+  "f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)"
+by (metis analytic_on_def singleton_iff)
+
+lemma analytic_at:
+    "f analytic_on {z} \<longleftrightarrow> (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s)"
+by (metis analytic_on_holomorphic empty_subsetI insert_subset)
+
+lemma analytic_on_analytic_at:
+    "f analytic_on s \<longleftrightarrow> (\<forall>z \<in> s. f analytic_on {z})"
+by (metis analytic_at_ball analytic_on_def)
+
+lemma analytic_at_two:
+  "f analytic_on {z} \<and> g analytic_on {z} \<longleftrightarrow>
+   (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s \<and> g holomorphic_on s)"
+  (is "?lhs = ?rhs")
+proof 
+  assume ?lhs
+  then obtain s t 
+    where st: "open s" "z \<in> s" "f holomorphic_on s"
+              "open t" "z \<in> t" "g holomorphic_on t"
+    by (auto simp: analytic_at)
+  show ?rhs
+    apply (rule_tac x="s \<inter> t" in exI)
+    using st
+    apply (auto simp: Diff_subset holomorphic_on_subset)
+    done
+next
+  assume ?rhs 
+  then show ?lhs
+    by (force simp add: analytic_at)
+qed
+
+subsection{*Combining theorems for derivative with ``analytic at'' hypotheses*}
+
+lemma 
+  assumes "f analytic_on {z}" "g analytic_on {z}"
+  shows complex_derivative_add_at: "DD (\<lambda>w. f w + g w) z = DD f z + DD g z"
+    and complex_derivative_diff_at: "DD (\<lambda>w. f w - g w) z = DD f z - DD g z"
+    and complex_derivative_mult_at: "DD (\<lambda>w. f w * g w) z =
+           f z * DD g z + DD f z * g z"
+proof -
+  obtain s where s: "open s" "z \<in> s" "f holomorphic_on s" "g holomorphic_on s"
+    using assms by (metis analytic_at_two)
+  show "DD (\<lambda>w. f w + g w) z = DD f z + DD g z"
+    apply (rule DERIV_imp_DD [OF DERIV_add])
+    using s
+    apply (auto simp: holomorphic_on_open complex_differentiable_def DD_iff_derivative_complex_differentiable)
+    done
+  show "DD (\<lambda>w. f w - g w) z = DD f z - DD g z"
+    apply (rule DERIV_imp_DD [OF DERIV_diff])
+    using s
+    apply (auto simp: holomorphic_on_open complex_differentiable_def DD_iff_derivative_complex_differentiable)
+    done
+  show "DD (\<lambda>w. f w * g w) z = f z * DD g z + DD f z * g z"
+    apply (rule DERIV_imp_DD [OF DERIV_mult'])
+    using s
+    apply (auto simp: holomorphic_on_open complex_differentiable_def DD_iff_derivative_complex_differentiable)
+    done
+qed
+
+lemma complex_derivative_cmult_at:
+  "f analytic_on {z} \<Longrightarrow>  DD (\<lambda>w. c * f w) z = c * DD f z"
+by (auto simp: complex_derivative_mult_at complex_derivative_const analytic_on_const)
+
+lemma complex_derivative_cmult_right_at:
+  "f analytic_on {z} \<Longrightarrow>  DD (\<lambda>w. f w * c) z = DD f z * c"
+by (auto simp: complex_derivative_mult_at complex_derivative_const analytic_on_const)
+
+text{*A composition lemma for functions of mixed type*}
+lemma has_vector_derivative_real_complex:
+  fixes f :: "complex \<Rightarrow> complex"
+  assumes "DERIV f (of_real a) :> f'"
+  shows "((\<lambda>x. f (of_real x)) has_vector_derivative f') (at a)"
+  using diff_chain_at [OF has_derivative_ident, of f "op * f'" a] assms
+  unfolding has_field_derivative_def has_vector_derivative_def o_def
+  by (auto simp: mult_ac scaleR_conv_of_real)
+
+subsection{*Complex differentiation of sequences and series*}
+
+lemma has_complex_derivative_sequence:
+  fixes s :: "complex set"
+  assumes cvs: "convex s"
+      and df:  "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)"
+      and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s \<longrightarrow> norm (f' n x - g' x) \<le> e"
+      and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) ---> l) sequentially"
+    shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) ---> g x) sequentially \<and> 
+                       (g has_field_derivative (g' x)) (at x within s)"
+proof -
+  from assms obtain x l where x: "x \<in> s" and tf: "((\<lambda>n. f n x) ---> l) sequentially"
+    by blast
+  { fix e::real assume e: "e > 0"
+    then obtain N where N: "\<forall>n\<ge>N. \<forall>x. x \<in> s \<longrightarrow> cmod (f' n x - g' x) \<le> e"
+      by (metis conv)    
+    have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
+    proof (rule exI [of _ N], clarify)
+      fix n y h
+      assume "N \<le> n" "y \<in> s"
+      then have "cmod (f' n y - g' y) \<le> e"
+        by (metis N)
+      then have "cmod h * cmod (f' n y - g' y) \<le> cmod h * e"
+        by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)
+      then show "cmod (f' n y * h - g' y * h) \<le> e * cmod h"
+        by (simp add: norm_mult [symmetric] field_simps)
+    qed
+  } note ** = this
+  show ?thesis
+  unfolding has_field_derivative_def
+  proof (rule has_derivative_sequence [OF cvs _ _ x])
+    show "\<forall>n. \<forall>x\<in>s. (f n has_derivative (op * (f' n x))) (at x within s)"
+      by (metis has_field_derivative_def df)
+  next show "(\<lambda>n. f n x) ----> l"
+    by (rule tf)
+  next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
+    by (blast intro: **)
+  qed
+qed
+
+
+lemma has_complex_derivative_series:
+  fixes s :: "complex set"
+  assumes cvs: "convex s"
+      and df:  "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)"
+      and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s 
+                \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e"
+      and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) sums l)"
+    shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((g has_field_derivative g' x) (at x within s))"
+proof -
+  from assms obtain x l where x: "x \<in> s" and sf: "((\<lambda>n. f n x) sums l)"
+    by blast
+  { fix e::real assume e: "e > 0"
+    then obtain N where N: "\<forall>n x. n \<ge> N \<longrightarrow> x \<in> s 
+            \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e"
+      by (metis conv)    
+    have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"
+    proof (rule exI [of _ N], clarify)
+      fix n y h
+      assume "N \<le> n" "y \<in> s"
+      then have "cmod ((\<Sum>i<n. f' i y) - g' y) \<le> e"
+        by (metis N)
+      then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e"
+        by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)
+      then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h"
+        by (simp add: norm_mult [symmetric] field_simps setsum_right_distrib)
+    qed
+  } note ** = this
+  show ?thesis
+  unfolding has_field_derivative_def
+  proof (rule has_derivative_series [OF cvs _ _ x])
+    fix n x
+    assume "x \<in> s"
+    then show "((f n) has_derivative (\<lambda>z. z * f' n x)) (at x within s)"
+      by (metis df has_field_derivative_def mult_commute_abs)
+  next show " ((\<lambda>n. f n x) sums l)"
+    by (rule sf)
+  next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"
+    by (blast intro: **)
+  qed
+qed
+
+subsection{*Bound theorem*}
+
+lemma complex_differentiable_bound:
+  fixes s :: "complex set"
+  assumes cvs: "convex s"
+      and df:  "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z within s)"
+      and dn:  "\<And>z. z \<in> s \<Longrightarrow> norm (f' z) \<le> B"
+      and "x \<in> s"  "y \<in> s"
+    shows "norm(f x - f y) \<le> B * norm(x - y)"
+  apply (rule differentiable_bound [OF cvs])
+  using assms
+  apply (auto simp: has_field_derivative_def Ball_def onorm_def)
+  apply (rule cSUP_least)
+  apply (auto simp: norm_mult)
+  apply (metis norm_one)
+  done
+
+subsection{*Inverse function theorem for complex derivatives.*}
+
+lemma has_complex_derivative_inverse_basic:
+  fixes f :: "complex \<Rightarrow> complex"
+  shows "DERIV f (g y) :> f' \<Longrightarrow>
+        f' \<noteq> 0 \<Longrightarrow>
+        continuous (at y) g \<Longrightarrow>
+        open t \<Longrightarrow>
+        y \<in> t \<Longrightarrow>
+        (\<And>z. z \<in> t \<Longrightarrow> f (g z) = z)
+        \<Longrightarrow> DERIV g y :> inverse (f')"
+  unfolding has_field_derivative_def
+  apply (rule has_derivative_inverse_basic)
+  apply (auto simp:  bounded_linear_mult_right)
+  done
+
+(*Used only once, in Multivariate/cauchy.ml. *)
+lemma has_complex_derivative_inverse_strong:
+  fixes f :: "complex \<Rightarrow> complex"
+  shows "DERIV f x :> f' \<Longrightarrow>
+         f' \<noteq> 0 \<Longrightarrow>
+         open s \<Longrightarrow>
+         x \<in> s \<Longrightarrow>
+         continuous_on s f \<Longrightarrow>
+         (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z)
+         \<Longrightarrow> DERIV g (f x) :> inverse (f')"
+  unfolding has_field_derivative_def
+  apply (rule has_derivative_inverse_strong [of s x f g ])
+  using assms 
+  by auto
+
+lemma has_complex_derivative_inverse_strong_x:
+  fixes f :: "complex \<Rightarrow> complex"
+  shows  "DERIV f (g y) :> f' \<Longrightarrow>
+          f' \<noteq> 0 \<Longrightarrow>
+          open s \<Longrightarrow>
+          continuous_on s f \<Longrightarrow>
+          g y \<in> s \<Longrightarrow> f(g y) = y \<Longrightarrow>
+          (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z)
+          \<Longrightarrow> DERIV g y :> inverse (f')"
+  unfolding has_field_derivative_def
+  apply (rule has_derivative_inverse_strong_x [of s g y f])
+  using assms 
+  by auto
+
+subsection{*Further useful properties of complex conjugation*}
+
+lemma lim_cnj: "((\<lambda>x. cnj(f x)) ---> cnj l) F \<longleftrightarrow> (f ---> l) F"
+  by (simp add: tendsto_iff dist_complex_def Complex.complex_cnj_diff [symmetric])
+
+lemma sums_cnj: "((\<lambda>x. cnj(f x)) sums cnj l) \<longleftrightarrow> (f sums l)"
+  by (simp add: sums_def lim_cnj cnj_setsum [symmetric])
+
+lemma continuous_within_cnj: "continuous (at z within s) cnj"
+by (metis bounded_linear_cnj linear_continuous_within)
+
+lemma continuous_on_cnj: "continuous_on s cnj"
+by (metis bounded_linear_cnj linear_continuous_on)
+
+subsection{*Some limit theorems about real part of real series etc.*}
+
+lemma real_lim:
+  fixes l::complex
+  assumes "(f ---> l) F" and " ~(trivial_limit F)" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>"
+  shows  "l \<in> \<real>"
+proof -
+  have 1: "((\<lambda>i. Im (f i)) ---> Im l) F"
+    by (metis assms(1) tendsto_Im) 
+  then have  "((\<lambda>i. Im (f i)) ---> 0) F"
+    by (smt2 Lim_eventually assms(3) assms(4) complex_is_Real_iff eventually_elim2)
+  then show ?thesis using 1
+    by (metis 1 assms(2) complex_is_Real_iff tendsto_unique) 
+qed
+
+lemma real_lim_sequentially:
+  fixes l::complex
+  shows "(f ---> l) sequentially \<Longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"
+by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially)
+
+lemma real_series: 
+  fixes l::complex
+  shows "f sums l \<Longrightarrow> (\<And>n. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"
+unfolding sums_def
+by (metis real_lim_sequentially setsum_in_Reals)
+
+
+lemma Lim_null_comparison_Re:
+   "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F \<Longrightarrow>  (g ---> 0) F \<Longrightarrow> (f ---> 0) F"
+  by (metis Lim_null_comparison complex_Re_zero tendsto_Re)
+
+
+lemma norm_setsum_bound:
+  fixes n::nat
+  shows" \<lbrakk>\<And>n. N \<le> n \<Longrightarrow> norm (f n) \<le> g n; N \<le> m\<rbrakk>
+       \<Longrightarrow> norm (setsum f {m..<n}) \<le> setsum g {m..<n}"
+apply (induct n, auto)
+by (metis dual_order.trans le_cases le_neq_implies_less norm_triangle_mono)
+
+(*Replaces the one in Series*)
+lemma summable_comparison_test:
+  fixes f:: "nat \<Rightarrow> 'a::banach" 
+  shows "summable g \<Longrightarrow> (\<And>n. n \<ge> N \<Longrightarrow> norm(f n) \<le> g n) \<Longrightarrow> summable f"
+apply (auto simp: Series.summable_Cauchy)  
+apply (drule_tac x = e in spec, auto)
+apply (rule_tac x="max N Na" in exI, auto)
+by (smt2 norm_setsum_bound)
+
+(*MOVE TO Finite_Cartesian_Product*)
+lemma sums_vec_nth :
+  assumes "f sums a"
+  shows "(\<lambda>x. f x $ i) sums a $ i"
+using assms unfolding sums_def
+by (auto dest: tendsto_vec_nth [where i=i])
+
+lemma summable_vec_nth :
+  assumes "summable f"
+  shows "summable (\<lambda>x. f x $ i)"
+using assms unfolding summable_def
+by (blast intro: sums_vec_nth)
+
+(*REPLACE THE ONES IN COMPLEX.THY*)
+lemma Re_setsum: "Re(setsum f s) = setsum (%x. Re(f x)) s"
+apply (cases "finite s")
+  by (induct s rule: finite_induct) auto
+
+lemma Im_setsum: "Im(setsum f s) = setsum (%x. Im(f x)) s"
+apply (cases "finite s")
+  by (induct s rule: finite_induct) auto
+
+lemma of_real_setsum: "of_real (setsum f s) = setsum (%x. of_real(f x)) s"
+apply (cases "finite s")
+  by (induct s rule: finite_induct) auto
+
+lemma of_real_setprod: "of_real (setprod f s) = setprod (%x. of_real(f x)) s"
+apply (cases "finite s")
+  by (induct s rule: finite_induct) auto
+
+lemma sums_Re:
+  assumes "f sums a"
+  shows "(\<lambda>x. Re (f x)) sums Re a"
+using assms 
+by (auto simp: sums_def Re_setsum [symmetric] isCont_tendsto_compose [of a Re])
+
+lemma sums_Im:
+  assumes "f sums a"
+  shows "(\<lambda>x. Im (f x)) sums Im a"
+using assms 
+by (auto simp: sums_def Im_setsum [symmetric] isCont_tendsto_compose [of a Im])
+
+lemma summable_Re:
+  assumes "summable f"
+  shows "summable (\<lambda>x. Re (f x))"
+using assms unfolding summable_def
+by (blast intro: sums_Re)
+
+lemma summable_Im:
+  assumes "summable f"
+  shows "summable (\<lambda>x. Im (f x))"
+using assms unfolding summable_def
+by (blast intro: sums_Im)
+
+lemma series_comparison_complex:
+  fixes f:: "nat \<Rightarrow> 'a::banach"
+  assumes sg: "summable g"
+     and "\<And>n. g n \<in> \<real>" "\<And>n. Re (g n) \<ge> 0"
+     and fg: "\<And>n. n \<ge> N \<Longrightarrow> norm(f n) \<le> norm(g n)"
+  shows "summable f"
+proof -
+  have g: "\<And>n. cmod (g n) = Re (g n)" using assms
+    by (metis abs_of_nonneg in_Reals_norm)
+  show ?thesis
+    apply (rule summable_comparison_test [where g = "\<lambda>n. norm (g n)" and N=N])
+    using sg
+    apply (auto simp: summable_def)
+    apply (rule_tac x="Re s" in exI)
+    apply (auto simp: g sums_Re)
+    apply (metis fg g)
+    done
+qed
+
+lemma summable_complex_of_real [simp]:
+  "summable (\<lambda>n. complex_of_real (f n)) = summable f"
+apply (auto simp: Series.summable_Cauchy)  
+apply (drule_tac x = e in spec, auto)
+apply (rule_tac x=N in exI)
+apply (auto simp: of_real_setsum [symmetric])
+done
+
+
+(* ------------------------------------------------------------------------- *)
+(* A kind of complex Taylor theorem.                                         *)
+(* ------------------------------------------------------------------------- *)
+
+
+lemma setsum_Suc_reindex:
+  fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
+    shows  "setsum f {0..n} = f 0 - f (Suc n) + setsum (\<lambda>i. f (Suc i)) {0..n}"
+by (induct n) auto
+
+(*REPLACE THE REAL VERSIONS*)
+lemma mult_left_cancel:
+  fixes c:: "'a::ring_no_zero_divisors"
+  shows "c \<noteq> 0 \<Longrightarrow> (c*a=c*b) = (a=b)"
+by simp 
+
+lemma mult_right_cancel:
+  fixes c:: "'a::ring_no_zero_divisors"
+  shows "c \<noteq> 0 \<Longrightarrow> (a*c=b*c) = (a=b)"
+by simp 
+
+lemma complex_taylor:
+  assumes s: "convex s" 
+      and f: "\<And>i x. x \<in> s \<Longrightarrow> i \<le> n \<Longrightarrow> (f i has_field_derivative f (Suc i) x) (at x within s)"
+      and B: "\<And>x. x \<in> s \<Longrightarrow> cmod (f (Suc n) x) \<le> B"
+      and w: "w \<in> s"
+      and z: "z \<in> s"
+    shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / of_nat (fact i)))
+          \<le> B * cmod(z - w)^(Suc n) / fact n"
+proof -
+  have wzs: "closed_segment w z \<subseteq> s" using assms
+    by (metis convex_contains_segment)
+  { fix u
+    assume "u \<in> closed_segment w z"
+    then have "u \<in> s"
+      by (metis wzs subsetD)
+    have "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / of_nat (fact i) +
+                      f (Suc i) u * (z-u)^i / of_nat (fact i)) = 
+              f (Suc n) u * (z-u) ^ n / of_nat (fact n)"
+    proof (induction n)
+      case 0 show ?case by simp
+    next
+      case (Suc n)
+      have "(\<Sum>i\<le>Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / of_nat (fact i) +
+                             f (Suc i) u * (z-u) ^ i / of_nat (fact i)) =  
+           f (Suc n) u * (z-u) ^ n / of_nat (fact n) +
+           f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / of_nat (fact (Suc n)) -
+           f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / of_nat (fact (Suc n))"
+        using Suc by simp
+      also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / of_nat (fact (Suc n))"
+      proof -
+        have "of_nat(fact(Suc n)) *
+             (f(Suc n) u *(z-u) ^ n / of_nat(fact n) +
+               f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / of_nat(fact(Suc n)) -
+               f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / of_nat(fact(Suc n))) =
+            (of_nat(fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / of_nat(fact n) +
+            (of_nat(fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / of_nat(fact(Suc n))) -
+            (of_nat(fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / of_nat(fact(Suc n))"
+          by (simp add: algebra_simps del: fact_Suc)
+        also have "... =
+                   (of_nat (fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / of_nat (fact n) +
+                   (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -
+                   (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"
+          by (simp del: fact_Suc)
+        also have "... = 
+                   (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) +
+                   (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -
+                   (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"
+          by (simp only: fact_Suc of_nat_mult mult_ac) simp
+        also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)"
+          by (simp add: algebra_simps)
+        finally show ?thesis
+        by (simp add: mult_left_cancel [where c = "of_nat (fact (Suc n))", THEN iffD1] del: fact_Suc)
+      qed
+      finally show ?case .
+    qed
+    then have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / of_nat (fact i))) 
+                has_field_derivative f (Suc n) u * (z-u) ^ n / of_nat (fact n))
+               (at u within s)"
+      apply (intro DERIV_intros DERIV_power[THEN DERIV_cong])
+      apply (blast intro: assms `u \<in> s`)
+      apply (rule refl)+
+      apply (auto simp: field_simps)
+      done
+  } note sum_deriv = this
+  { fix u
+    assume u: "u \<in> closed_segment w z"
+    then have us: "u \<in> s"
+      by (metis wzs subsetD)
+    have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> cmod (f (Suc n) u) * cmod (u - z) ^ n"
+      by (metis norm_minus_commute order_refl)
+    also have "... \<le> cmod (f (Suc n) u) * cmod (z - w) ^ n"
+      by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u])
+    also have "... \<le> B * cmod (z - w) ^ n"
+      by (metis norm_ge_zero zero_le_power mult_right_mono  B [OF us])
+    finally have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> B * cmod (z - w) ^ n" .
+  } note cmod_bound = this
+  have "(\<Sum>i\<le>n. f i z * (z - z) ^ i / of_nat (fact i)) = (\<Sum>i\<le>n. (f i z / of_nat (fact i)) * 0 ^ i)"
+    by simp
+  also have "\<dots> = f 0 z / of_nat (fact 0)"
+    by (subst setsum_zero_power) simp
+  finally have "cmod (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / of_nat (fact i))) 
+            \<le> cmod ((\<Sum>i\<le>n. f i w * (z - w) ^ i / of_nat (fact i)) -
+                    (\<Sum>i\<le>n. f i z * (z - z) ^ i / of_nat (fact i)))"
+    by (simp add: norm_minus_commute)
+  also have "... \<le> B * cmod (z - w) ^ n / real_of_nat (fact n) * cmod (w - z)"
+    apply (rule complex_differentiable_bound 
+      [where f' = "\<lambda>w. f (Suc n) w * (z - w)^n / of_nat(fact n)"
+         and s = "closed_segment w z", OF convex_segment])
+    apply (auto simp: ends_in_segment real_of_nat_def DERIV_subset [OF sum_deriv wzs]
+                  norm_divide norm_mult norm_power divide_le_cancel cmod_bound)
+    done
+  also have "...  \<le> B * cmod (z - w) ^ Suc n / real (fact n)"
+    by (simp add: algebra_simps norm_minus_commute real_of_nat_def)
+  finally show ?thesis .
+qed
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/PolyRoots.thy	Wed Mar 19 14:54:45 2014 +0000
@@ -0,0 +1,296 @@
+header {* polynomial functions: extremal behaviour and root counts *}
+
+(*  Author: John Harrison and Valentina Bruno
+    Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson
+*)
+
+theory PolyRoots
+imports Complex_Main
+
+begin
+
+subsection{*Geometric progressions*}
+
+lemma setsum_gp_basic:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  shows "(1 - x) * (\<Sum>i\<le>n. x^i) = 1 - x^Suc n"
+  by (simp only: one_diff_power_eq [of "Suc n" x] lessThan_Suc_atMost)
+
+lemma setsum_gp0:
+ fixes x :: "'a::{comm_ring,division_ring_inverse_zero}"
+ shows   "(\<Sum>i\<le>n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))"
+using setsum_gp_basic[of x n]
+apply (simp add: real_of_nat_def)
+by (metis eq_iff_diff_eq_0 mult_commute nonzero_eq_divide_eq)
+
+lemma setsum_power_shift:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  assumes "m \<le> n"
+  shows "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i\<le>n-m. x^i)"
+proof -
+  have "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i=m..n. x^(i-m))"
+    by (simp add: setsum_right_distrib power_add [symmetric])
+  also have "... = x^m * (\<Sum>i\<le>n-m. x^i)"
+    apply (subst setsum_reindex_cong [where f = "%i. i+m" and A = "{..n-m}"])
+    apply (auto simp: image_def)
+    apply (rule_tac x="x-m" in bexI, auto)
+    by (metis assms ordered_cancel_comm_monoid_diff_class.le_diff_conv2)
+  finally show ?thesis .
+qed
+
+lemma setsum_gp_multiplied:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  assumes "m \<le> n"
+  shows "(1 - x) * (\<Sum>i=m..n. x^i) = x^m - x^Suc n"
+proof -
+  have  "(1 - x) * (\<Sum>i=m..n. x^i) = x^m * (1 - x) * (\<Sum>i\<le>n-m. x^i)"
+    by (metis ab_semigroup_mult_class.mult_ac(1) assms mult_commute setsum_power_shift)
+  also have "... =x^m * (1 - x^Suc(n-m))"
+    by (metis ab_semigroup_mult_class.mult_ac(1) setsum_gp_basic)
+  also have "... = x^m - x^Suc n"
+    using assms
+    by (simp add: algebra_simps) (metis le_add_diff_inverse power_add)
+  finally show ?thesis .
+qed
+
+lemma setsum_gp:
+  fixes x :: "'a::{comm_ring,division_ring_inverse_zero}"
+  shows   "(\<Sum>i=m..n. x^i) =
+               (if n < m then 0
+                else if x = 1 then of_nat((n + 1) - m)
+                else (x^m - x^Suc n) / (1 - x))"
+using setsum_gp_multiplied [of m n x] 
+apply (auto simp: real_of_nat_def)
+by (metis eq_iff_diff_eq_0 mult_commute nonzero_divide_eq_eq)
+
+lemma setsum_gp_offset:
+  fixes x :: "'a::{comm_ring,division_ring_inverse_zero}"
+  shows   "(\<Sum>i=m..m+n. x^i) =
+       (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
+  using setsum_gp [of x m "m+n"]
+  by (auto simp: power_add algebra_simps)
+
+subsection{*Basics about polynomial functions: extremal behaviour and root counts.*}
+
+lemma sub_polyfun:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  shows   "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = 
+           (x - y) * (\<Sum>j<n. \<Sum>k= Suc j..n. a k * y^(k - Suc j) * x^j)"
+proof -
+  have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = 
+        (\<Sum>i\<le>n. a i * (x^i - y^i))"
+    by (simp add: algebra_simps setsum_subtractf [symmetric])
+  also have "... = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"
+    by (simp add: power_diff_sumr2 mult_ac)
+  also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))"
+    by (simp add: setsum_right_distrib mult_ac)
+  also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))"
+    by (simp add: nested_setsum_swap')
+  finally show ?thesis .
+qed
+
+lemma sub_polyfun_alt:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  shows   "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = 
+           (x - y) * (\<Sum>j<n. \<Sum>k<n-j. a (j+k+1) * y^k * x^j)"
+proof -
+  { fix j
+    assume "j < n"
+    have "(\<Sum>k = Suc j..n. a k * y^(k - Suc j) * x^j) =
+          (\<Sum>k <n - j. a (Suc (j + k)) * y^k * x^j)"
+      apply (rule_tac f = "\<lambda>i. Suc j + i" in setsum_reindex_cong)
+      apply (auto simp: inj_on_def image_def atLeastLessThan_def lessThan_def)
+      apply (metis Suc_le_eq diff_add_inverse diff_less_mono le_add1 less_imp_Suc_add)
+      done }
+  then show ?thesis
+    by (simp add: sub_polyfun)
+qed
+
+lemma polyfun_linear_factor:
+  fixes a :: "'a::{comm_ring,monoid_mult}"
+  shows  "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = 
+                  (z-a) * (\<Sum>i<n. b i * z^i) + (\<Sum>i\<le>n. c i * a^i)"
+proof -
+  { fix z
+    have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) = 
+          (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)"
+      by (simp add: sub_polyfun setsum_left_distrib)
+    then have "(\<Sum>i\<le>n. c i * z^i) = 
+          (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)
+          + (\<Sum>i\<le>n. c i * a^i)"
+      by (simp add: algebra_simps) }
+  then show ?thesis
+    by (intro exI allI) 
+qed
+
+lemma polyfun_linear_factor_root:
+  fixes a :: "'a::{comm_ring,monoid_mult}"
+  assumes "(\<Sum>i\<le>n. c i * a^i) = 0"
+  shows  "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = (z-a) * (\<Sum>i<n. b i * z^i)"
+  using polyfun_linear_factor [of c n a] assms
+  by simp
+
+lemma adhoc_norm_triangle: "a + norm(y) \<le> b ==> norm(x) \<le> a ==> norm(x + y) \<le> b"
+  by (metis norm_triangle_mono order.trans order_refl)
+
+lemma polyfun_extremal_lemma:
+  fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
+  assumes "e > 0"
+    shows "\<exists>M. \<forall>z. M \<le> norm z \<longrightarrow> norm(\<Sum>i\<le>n. c i * z^i) \<le> e * norm(z) ^ Suc n"
+proof (induction n)
+  case 0
+  show ?case 
+    by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult_commute pos_divide_le_eq assms)
+next
+  case (Suc n)
+  then obtain M where M: "\<forall>z. M \<le> norm z \<longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" ..
+  show ?case
+  proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify)
+    fix z::'a
+    assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \<le> norm z"
+    then have norm1: "0 < norm z" "M \<le> norm z" "(e + norm (c (Suc n))) / e \<le> norm z"
+      by auto
+    then have norm2: "(e + norm (c (Suc n))) \<le> e * norm z"  "(norm z * norm z ^ n) > 0"
+      apply (metis assms less_divide_eq mult_commute not_le) 
+      using norm1 apply (metis mult_pos_pos zero_less_power)
+      done
+    have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) =
+          (e + norm (c (Suc n))) * (norm z * norm z ^ n)"
+      by (simp add: norm_mult norm_power algebra_simps)
+    also have "... \<le> (e * norm z) * (norm z * norm z ^ n)"
+      using norm2 by (metis real_mult_le_cancel_iff1) 
+    also have "... = e * (norm z * (norm z * norm z ^ n))"
+      by (simp add: algebra_simps)
+    finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n))
+                  \<le> e * (norm z * (norm z * norm z ^ n))" .
+    then show "norm (\<Sum>i\<le>Suc n. c i * z^i) \<le> e * norm z ^ Suc (Suc n)" using M norm1
+      by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle)
+    qed
+qed
+
+lemma norm_lemma_xy: "\<lbrakk>abs b + 1 \<le> norm(y) - a; norm(x) \<le> a\<rbrakk> \<Longrightarrow> b \<le> norm(x + y)"
+  by (metis abs_add_one_not_less_self add_commute diff_le_eq dual_order.trans le_less_linear 
+         norm_diff_ineq)
+
+lemma polyfun_extremal:
+  fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
+  assumes "\<exists>k. k \<noteq> 0 \<and> k \<le> n \<and> c k \<noteq> 0"
+    shows "eventually (\<lambda>z. norm(\<Sum>i\<le>n. c i * z^i) \<ge> B) at_infinity"
+using assms
+proof (induction n)
+  case 0 then show ?case
+    by simp
+next
+  case (Suc n)
+  show ?case
+  proof (cases "c (Suc n) = 0")
+    case True
+    with Suc show ?thesis
+      by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq)
+  next
+    case False
+    with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n]
+    obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow> 
+               norm (\<Sum>i\<le>n. c i * z^i) \<le> norm (c (Suc n)) / 2 * norm z ^ Suc n"
+      by auto
+    show ?thesis
+    unfolding eventually_at_infinity
+    proof (rule exI [where x="max M (max 1 ((abs B + 1) / (norm (c (Suc n)) / 2)))"], clarsimp)
+      fix z::'a
+      assume les: "M \<le> norm z"  "1 \<le> norm z"  "(\<bar>B\<bar> * 2 + 2) / norm (c (Suc n)) \<le> norm z"
+      then have "\<bar>B\<bar> * 2 + 2 \<le> norm z * norm (c (Suc n))"
+        by (metis False pos_divide_le_eq zero_less_norm_iff)
+      then have "\<bar>B\<bar> * 2 + 2 \<le> norm z ^ (Suc n) * norm (c (Suc n))" 
+        by (metis `1 \<le> norm z` order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc)
+      then show "B \<le> norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les
+        apply auto
+        apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"])
+        apply (simp_all add: norm_mult norm_power)
+        done
+    qed
+  qed
+qed
+
+lemma polyfun_rootbound:
+ fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+ assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"
+   shows "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<and> card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
+using assms
+proof (induction n arbitrary: c)
+ case (Suc n) show ?case
+ proof (cases "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = {}")
+   case False
+   then obtain a where a: "(\<Sum>i\<le>Suc n. c i * a^i) = 0"
+     by auto
+   from polyfun_linear_factor_root [OF this]
+   obtain b where "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i< Suc n. b i * z^i)"
+     by auto
+   then have b: "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i\<le>n. b i * z^i)"
+     by (metis lessThan_Suc_atMost)
+   then have ins_ab: "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = insert a {z. (\<Sum>i\<le>n. b i * z^i) = 0}"
+     by auto
+   have c0: "c 0 = - (a * b 0)" using  b [of 0]
+     by simp
+   then have extr_prem: "~ (\<exists>k\<le>n. b k \<noteq> 0) \<Longrightarrow> \<exists>k. k \<noteq> 0 \<and> k \<le> Suc n \<and> c k \<noteq> 0"
+     by (metis Suc.prems le0 minus_zero mult_zero_right)
+   have "\<exists>k\<le>n. b k \<noteq> 0" 
+     apply (rule ccontr)
+     using polyfun_extremal [OF extr_prem, of 1]
+     apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc)
+     apply (drule_tac x="of_real ba" in spec, simp)
+     done
+   then show ?thesis using Suc.IH [of b] ins_ab
+     by (auto simp: card_insert_if)
+   qed simp
+qed simp
+
+corollary
+  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+  assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"
+    shows polyfun_rootbound_finite: "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
+      and polyfun_rootbound_card:   "card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
+using polyfun_rootbound [OF assms] by auto
+
+lemma polyfun_finite_roots:
+  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+    shows  "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<longleftrightarrow> (\<exists>k. k \<le> n \<and> c k \<noteq> 0)"
+proof (cases " \<exists>k\<le>n. c k \<noteq> 0")
+  case True then show ?thesis 
+    by (blast intro: polyfun_rootbound_finite)
+next
+  case False then show ?thesis 
+    by (auto simp: infinite_UNIV_char_0)
+qed
+
+lemma polyfun_eq_0:
+  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+    shows  "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0) \<longleftrightarrow> (\<forall>k. k \<le> n \<longrightarrow> c k = 0)"
+proof (cases "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0)")
+  case True
+  then have "~ finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
+    by (simp add: infinite_UNIV_char_0)
+  with True show ?thesis
+    by (metis (poly_guards_query) polyfun_rootbound_finite)
+next
+  case False
+  then show ?thesis
+    by auto
+qed
+
+lemma polyfun_eq_const:
+  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+    shows  "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"
+proof -
+  {fix z
+    have "(\<Sum>i\<le>n. c i * z^i) = (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) + k"
+      by (induct n) auto
+  } then
+  have "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> (\<forall>z. (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) = 0)"
+    by auto
+  also have "... \<longleftrightarrow>  c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"
+    by (auto simp: polyfun_eq_0)
+  finally show ?thesis .
+qed
+
+end
+
--- a/src/HOL/ROOT	Tue Mar 18 22:11:46 2014 +0100
+++ b/src/HOL/ROOT	Wed Mar 19 14:54:45 2014 +0000
@@ -677,6 +677,8 @@
   theories
     Multivariate_Analysis
     Determinants
+    PolyRoots
+    Complex_Analysis_Basics
   files
     "document/root.tex"
 
--- a/src/HOL/Set_Interval.thy	Tue Mar 18 22:11:46 2014 +0100
+++ b/src/HOL/Set_Interval.thy	Wed Mar 19 14:54:45 2014 +0000
@@ -1476,6 +1476,16 @@
      "(\<Sum>i = 0..n. (\<Sum>j = 0..<i. a i j)) = (\<Sum>j = 0..<n. \<Sum>i = Suc j..n. a i j)"
   by (induction n) (auto simp: setsum_addf)
 
+lemma nested_setsum_swap':
+     "(\<Sum>i\<le>n. (\<Sum>j<i. a i j)) = (\<Sum>j<n. \<Sum>i = Suc j..n. a i j)"
+  by (induction n) (auto simp: setsum_addf)
+
+lemma setsum_zero_power [simp]:
+  fixes c :: "nat \<Rightarrow> 'a::division_ring"
+  shows "(\<Sum>i\<in>A. c i * 0^i) = (if finite A \<and> 0 \<in> A then c 0 else 0)"
+apply (cases "finite A")
+  by (induction A rule: finite_induct) auto
+
 
 subsection {* The formula for geometric sums *}