diff --git a/src/HOL/Analysis/Elementary_Normed_Spaces.thy b/src/HOL/Analysis/Elementary_Normed_Spaces.thy --- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy +++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy @@ -1,1755 +1,1789 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) section \Elementary Normed Vector Spaces\ theory Elementary_Normed_Spaces imports "HOL-Library.FuncSet" Elementary_Metric_Spaces Cartesian_Space Connected begin subsection \Orthogonal Transformation of Balls\ subsection\<^marker>\tag unimportant\ \Various Lemmas Combining Imports\ lemma open_sums: fixes T :: "('b::real_normed_vector) set" assumes "open S \ open T" shows "open (\x\ S. \y \ T. {x + y})" using assms proof assume S: "open S" show ?thesis proof (clarsimp simp: open_dist) fix x y assume "x \ S" "y \ T" with S obtain e where "e > 0" and e: "\x'. dist x' x < e \ x' \ S" by (auto simp: open_dist) then have "\z. dist z (x + y) < e \ \x\S. \y\T. z = x + y" by (metis \y \ T\ diff_add_cancel dist_add_cancel2) then show "\e>0. \z. dist z (x + y) < e \ (\x\S. \y\T. z = x + y)" using \0 < e\ \x \ S\ by blast qed next assume T: "open T" show ?thesis proof (clarsimp simp: open_dist) fix x y assume "x \ S" "y \ T" with T obtain e where "e > 0" and e: "\x'. dist x' y < e \ x' \ T" by (auto simp: open_dist) then have "\z. dist z (x + y) < e \ \x\S. \y\T. z = x + y" by (metis \x \ S\ add_diff_cancel_left' add_diff_eq diff_diff_add dist_norm) then show "\e>0. \z. dist z (x + y) < e \ (\x\S. \y\T. z = x + y)" using \0 < e\ \y \ T\ by blast qed qed lemma image_orthogonal_transformation_ball: fixes f :: "'a::euclidean_space \ 'a" assumes "orthogonal_transformation f" shows "f ` ball x r = ball (f x) r" proof (intro equalityI subsetI) fix y assume "y \ f ` ball x r" with assms show "y \ ball (f x) r" by (auto simp: orthogonal_transformation_isometry) next fix y assume y: "y \ ball (f x) r" then obtain z where z: "y = f z" using assms orthogonal_transformation_surj by blast with y assms show "y \ f ` ball x r" by (auto simp: orthogonal_transformation_isometry) qed lemma image_orthogonal_transformation_cball: fixes f :: "'a::euclidean_space \ 'a" assumes "orthogonal_transformation f" shows "f ` cball x r = cball (f x) r" proof (intro equalityI subsetI) fix y assume "y \ f ` cball x r" with assms show "y \ cball (f x) r" by (auto simp: orthogonal_transformation_isometry) next fix y assume y: "y \ cball (f x) r" then obtain z where z: "y = f z" using assms orthogonal_transformation_surj by blast with y assms show "y \ f ` cball x r" by (auto simp: orthogonal_transformation_isometry) qed subsection \Support\ definition (in monoid_add) support_on :: "'b set \ ('b \ 'a) \ 'b set" where "support_on S f = {x\S. f x \ 0}" lemma in_support_on: "x \ support_on S f \ x \ S \ f x \ 0" by (simp add: support_on_def) lemma support_on_simps[simp]: "support_on {} f = {}" "support_on (insert x S) f = (if f x = 0 then support_on S f else insert x (support_on S f))" "support_on (S \ T) f = support_on S f \ support_on T f" "support_on (S \ T) f = support_on S f \ support_on T f" "support_on (S - T) f = support_on S f - support_on T f" "support_on (f ` S) g = f ` (support_on S (g \ f))" unfolding support_on_def by auto lemma support_on_cong: "(\x. x \ S \ f x = 0 \ g x = 0) \ support_on S f = support_on S g" by (auto simp: support_on_def) lemma support_on_if: "a \ 0 \ support_on A (\x. if P x then a else 0) = {x\A. P x}" by (auto simp: support_on_def) lemma support_on_if_subset: "support_on A (\x. if P x then a else 0) \ {x \ A. P x}" by (auto simp: support_on_def) lemma finite_support[intro]: "finite S \ finite (support_on S f)" unfolding support_on_def by auto (* TODO: is supp_sum really needed? TODO: Generalize to Finite_Set.fold *) definition (in comm_monoid_add) supp_sum :: "('b \ 'a) \ 'b set \ 'a" where "supp_sum f S = (\x\support_on S f. f x)" lemma supp_sum_empty[simp]: "supp_sum f {} = 0" unfolding supp_sum_def by auto lemma supp_sum_insert[simp]: "finite (support_on S f) \ supp_sum f (insert x S) = (if x \ S then supp_sum f S else f x + supp_sum f S)" by (simp add: supp_sum_def in_support_on insert_absorb) lemma supp_sum_divide_distrib: "supp_sum f A / (r::'a::field) = supp_sum (\n. f n / r) A" by (cases "r = 0") (auto simp: supp_sum_def sum_divide_distrib intro!: sum.cong support_on_cong) subsection \Intervals\ lemma image_affinity_interval: fixes c :: "'a::ordered_real_vector" shows "((\x. m *\<^sub>R x + c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {m *\<^sub>R a + c .. m *\<^sub>R b + c} else {m *\<^sub>R b + c .. m *\<^sub>R a + c})" (is "?lhs = ?rhs") proof (cases "m=0") case True then show ?thesis by force next case False show ?thesis proof show "?lhs \ ?rhs" by (auto simp: scaleR_left_mono scaleR_left_mono_neg) show "?rhs \ ?lhs" proof (clarsimp, intro conjI impI subsetI) show "\0 \ m; a \ b; x \ {m *\<^sub>R a + c..m *\<^sub>R b + c}\ \ x \ (\x. m *\<^sub>R x + c) ` {a..b}" for x using False by (rule_tac x="inverse m *\<^sub>R (x-c)" in image_eqI) (auto simp: pos_le_divideR_eq pos_divideR_le_eq le_diff_eq diff_le_eq) show "\\ 0 \ m; a \ b; x \ {m *\<^sub>R b + c..m *\<^sub>R a + c}\ \ x \ (\x. m *\<^sub>R x + c) ` {a..b}" for x by (rule_tac x="inverse m *\<^sub>R (x-c)" in image_eqI) (auto simp add: neg_le_divideR_eq neg_divideR_le_eq le_diff_eq diff_le_eq) qed qed qed subsection \Limit Points\ lemma islimpt_ball: fixes x y :: "'a::{real_normed_vector,perfect_space}" shows "y islimpt ball x e \ 0 < e \ y \ cball x e" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof { assume "e \ 0" then have *: "ball x e = {}" using ball_eq_empty[of x e] by auto have False using \?lhs\ unfolding * using islimpt_EMPTY[of y] by auto } then show "e > 0" by (metis not_less) show "y \ cball x e" using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] ball_subset_cball[of x e] \?lhs\ unfolding closed_limpt by auto qed show ?lhs if ?rhs proof - from that have "e > 0" by auto { fix d :: real assume "d > 0" have "\x'\ball x e. x' \ y \ dist x' y < d" proof (cases "d \ dist x y") case True then show ?thesis proof (cases "x = y") case True then have False using \d \ dist x y\ \d>0\ by auto then show ?thesis by auto next case False have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))" unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[symmetric] by auto also have "\ = \- 1 + d / (2 * norm (x - y))\ * norm (x - y)" using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", symmetric, of "y - x"] unfolding scaleR_minus_left scaleR_one by (auto simp: norm_minus_commute) also have "\ = \- norm (x - y) + d / 2\" unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]] unfolding distrib_right using \x\y\ by auto also have "\ \ e - d/2" using \d \ dist x y\ and \d>0\ and \?rhs\ by (auto simp: dist_norm) finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \ ball x e" using \d>0\ by auto moreover have "(d / (2*dist y x)) *\<^sub>R (y - x) \ 0" using \x\y\[unfolded dist_nz] \d>0\ unfolding scaleR_eq_0_iff by (auto simp: dist_commute) moreover have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" using \0 < d\ by (fastforce simp: dist_norm) ultimately show ?thesis by (rule_tac x = "y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto qed next case False then have "d > dist x y" by auto show "\x' \ ball x e. x' \ y \ dist x' y < d" proof (cases "x = y") case True obtain z where z: "z \ y" "dist z y < min e d" using perfect_choose_dist[of "min e d" y] using \d > 0\ \e>0\ by auto show ?thesis by (metis True z dist_commute mem_ball min_less_iff_conj) next case False then show ?thesis using \d>0\ \d > dist x y\ \?rhs\ by force qed qed } then show ?thesis unfolding mem_cball islimpt_approachable mem_ball by auto qed qed lemma closure_ball_lemma: fixes x y :: "'a::real_normed_vector" assumes "x \ y" shows "y islimpt ball x (dist x y)" proof (rule islimptI) fix T assume "y \ T" "open T" then obtain r where "0 < r" "\z. dist z y < r \ z \ T" unfolding open_dist by fast \\choose point between @{term x} and @{term y}, within distance @{term r} of @{term y}.\ define k where "k = min 1 (r / (2 * dist x y))" define z where "z = y + scaleR k (x - y)" have z_def2: "z = x + scaleR (1 - k) (y - x)" unfolding z_def by (simp add: algebra_simps) have "dist z y < r" unfolding z_def k_def using \0 < r\ by (simp add: dist_norm min_def) then have "z \ T" using \\z. dist z y < r \ z \ T\ by simp have "dist x z < dist x y" using \0 < r\ assms by (simp add: z_def2 k_def dist_norm norm_minus_commute) then have "z \ ball x (dist x y)" by simp have "z \ y" unfolding z_def k_def using \x \ y\ \0 < r\ by (simp add: min_def) show "\z\ball x (dist x y). z \ T \ z \ y" using \z \ ball x (dist x y)\ \z \ T\ \z \ y\ by fast qed subsection \Balls and Spheres in Normed Spaces\ lemma mem_ball_0 [simp]: "x \ ball 0 e \ norm x < e" for x :: "'a::real_normed_vector" by simp lemma mem_cball_0 [simp]: "x \ cball 0 e \ norm x \ e" for x :: "'a::real_normed_vector" by simp lemma closure_ball [simp]: fixes x :: "'a::real_normed_vector" assumes "0 < e" shows "closure (ball x e) = cball x e" proof show "closure (ball x e) \ cball x e" using closed_cball closure_minimal by blast have "\y. dist x y < e \ dist x y = e \ y \ closure (ball x e)" by (metis Un_iff assms closure_ball_lemma closure_def dist_eq_0_iff mem_Collect_eq mem_ball) then show "cball x e \ closure (ball x e)" by force qed lemma mem_sphere_0 [simp]: "x \ sphere 0 e \ norm x = e" for x :: "'a::real_normed_vector" by simp (* In a trivial vector space, this fails for e = 0. *) lemma interior_cball [simp]: fixes x :: "'a::{real_normed_vector, perfect_space}" shows "interior (cball x e) = ball x e" proof (cases "e \ 0") case False note cs = this from cs have null: "ball x e = {}" using ball_empty[of e x] by auto moreover have "cball x e = {}" proof (rule equals0I) fix y assume "y \ cball x e" then show False by (metis ball_eq_empty null cs dist_eq_0_iff dist_le_zero_iff empty_subsetI mem_cball subset_antisym subset_ball) qed then have "interior (cball x e) = {}" using interior_empty by auto ultimately show ?thesis by blast next case True note cs = this have "ball x e \ cball x e" using ball_subset_cball by auto moreover { fix S y assume as: "S \ cball x e" "open S" "y\S" then obtain d where "d>0" and d: "\x'. dist x' y < d \ x' \ S" unfolding open_dist by blast then obtain xa where xa_y: "xa \ y" and xa: "dist xa y < d" using perfect_choose_dist [of d] by auto have "xa \ S" using d[THEN spec[where x = xa]] using xa by (auto simp: dist_commute) then have xa_cball: "xa \ cball x e" using as(1) by auto then have "y \ ball x e" proof (cases "x = y") case True then have "e > 0" using cs order.order_iff_strict xa_cball xa_y by fastforce then show "y \ ball x e" using \x = y \ by simp next case False have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" unfolding dist_norm using \d>0\ norm_ge_zero[of "y - x"] \x \ y\ by auto then have *: "y + (d / 2 / dist y x) *\<^sub>R (y - x) \ cball x e" using d as(1)[unfolded subset_eq] by blast have "y - x \ 0" using \x \ y\ by auto hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[symmetric] using \d>0\ by auto have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)" by (auto simp: dist_norm algebra_simps) also have "\ = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))" by (auto simp: algebra_simps) also have "\ = \1 + d / (2 * norm (y - x))\ * norm (y - x)" using ** by auto also have "\ = (dist y x) + d/2" using ** by (auto simp: distrib_right dist_norm) finally have "e \ dist x y +d/2" using *[unfolded mem_cball] by (auto simp: dist_commute) then show "y \ ball x e" unfolding mem_ball using \d>0\ by auto qed } then have "\S \ cball x e. open S \ S \ ball x e" by auto ultimately show ?thesis using interior_unique[of "ball x e" "cball x e"] using open_ball[of x e] by auto qed lemma frontier_ball [simp]: fixes a :: "'a::real_normed_vector" shows "0 < e \ frontier (ball a e) = sphere a e" by (force simp: frontier_def) lemma frontier_cball [simp]: fixes a :: "'a::{real_normed_vector, perfect_space}" shows "frontier (cball a e) = sphere a e" by (force simp: frontier_def) corollary compact_sphere [simp]: fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}" shows "compact (sphere a r)" using compact_frontier [of "cball a r"] by simp corollary bounded_sphere [simp]: fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}" shows "bounded (sphere a r)" by (simp add: compact_imp_bounded) corollary closed_sphere [simp]: fixes a :: "'a::{real_normed_vector,perfect_space,heine_borel}" shows "closed (sphere a r)" by (simp add: compact_imp_closed) lemma image_add_ball [simp]: fixes a :: "'a::real_normed_vector" shows "(+) b ` ball a r = ball (a+b) r" proof - { fix x :: 'a assume "dist (a + b) x < r" moreover have "b + (x - b) = x" by simp ultimately have "x \ (+) b ` ball a r" by (metis add.commute dist_add_cancel image_eqI mem_ball) } then show ?thesis by (auto simp: add.commute) qed lemma image_add_cball [simp]: fixes a :: "'a::real_normed_vector" shows "(+) b ` cball a r = cball (a+b) r" proof - have "\x. dist (a + b) x \ r \ \y\cball a r. x = b + y" by (metis (no_types) add.commute diff_add_cancel dist_add_cancel2 mem_cball) then show ?thesis by (force simp: add.commute) qed subsection\<^marker>\tag unimportant\ \Various Lemmas on Normed Algebras\ lemma closed_of_nat_image: "closed (of_nat ` A :: 'a::real_normed_algebra_1 set)" by (rule discrete_imp_closed[of 1]) (auto simp: dist_of_nat) lemma closed_of_int_image: "closed (of_int ` A :: 'a::real_normed_algebra_1 set)" by (rule discrete_imp_closed[of 1]) (auto simp: dist_of_int) lemma closed_Nats [simp]: "closed (\ :: 'a :: real_normed_algebra_1 set)" unfolding Nats_def by (rule closed_of_nat_image) lemma closed_Ints [simp]: "closed (\ :: 'a :: real_normed_algebra_1 set)" unfolding Ints_def by (rule closed_of_int_image) lemma closed_subset_Ints: fixes A :: "'a :: real_normed_algebra_1 set" assumes "A \ \" shows "closed A" proof (intro discrete_imp_closed[OF zero_less_one] ballI impI, goal_cases) case (1 x y) with assms have "x \ \" and "y \ \" by auto with \dist y x < 1\ show "y = x" by (auto elim!: Ints_cases simp: dist_of_int) qed subsection \Filters\ definition indirection :: "'a::real_normed_vector \ 'a \ 'a filter" (infixr "indirection" 70) where "a indirection v = at a within {b. \c\0. b - a = scaleR c v}" subsection \Trivial Limits\ lemma trivial_limit_at_infinity: "\ trivial_limit (at_infinity :: ('a::{real_normed_vector,perfect_space}) filter)" proof - obtain x::'a where "x \ 0" by (meson perfect_choose_dist zero_less_one) then have "b \ norm ((b / norm x) *\<^sub>R x)" for b by simp then show ?thesis unfolding trivial_limit_def eventually_at_infinity by blast qed lemma at_within_ball_bot_iff: fixes x y :: "'a::{real_normed_vector,perfect_space}" shows "at x within ball y r = bot \ (r=0 \ x \ cball y r)" unfolding trivial_limit_within by (metis (no_types) cball_empty equals0D islimpt_ball less_linear) subsection \Limits\ proposition Lim_at_infinity: "(f \ l) at_infinity \ (\e>0. \b. \x. norm x \ b \ dist (f x) l < e)" by (auto simp: tendsto_iff eventually_at_infinity) corollary Lim_at_infinityI [intro?]: assumes "\e. e > 0 \ \B. \x. norm x \ B \ dist (f x) l \ e" shows "(f \ l) at_infinity" proof - have "\e. e > 0 \ \B. \x. norm x \ B \ dist (f x) l < e" by (meson assms dense le_less_trans) then show ?thesis using Lim_at_infinity by blast qed lemma Lim_transform_within_set_eq: fixes a :: "'a::metric_space" and l :: "'b::metric_space" shows "eventually (\x. x \ S \ x \ T) (at a) \ ((f \ l) (at a within S) \ (f \ l) (at a within T))" by (force intro: Lim_transform_within_set elim: eventually_mono) lemma Lim_null: fixes f :: "'a \ 'b::real_normed_vector" shows "(f \ l) net \ ((\x. f(x) - l) \ 0) net" by (simp add: Lim dist_norm) lemma Lim_null_comparison: fixes f :: "'a \ 'b::real_normed_vector" assumes "eventually (\x. norm (f x) \ g x) net" "(g \ 0) net" shows "(f \ 0) net" using assms(2) proof (rule metric_tendsto_imp_tendsto) show "eventually (\x. dist (f x) 0 \ dist (g x) 0) net" using assms(1) by (rule eventually_mono) (simp add: dist_norm) qed lemma Lim_transform_bound: fixes f :: "'a \ 'b::real_normed_vector" and g :: "'a \ 'c::real_normed_vector" assumes "eventually (\n. norm (f n) \ norm (g n)) net" and "(g \ 0) net" shows "(f \ 0) net" using assms(1) tendsto_norm_zero [OF assms(2)] by (rule Lim_null_comparison) lemma lim_null_mult_right_bounded: fixes f :: "'a \ 'b::real_normed_div_algebra" assumes f: "(f \ 0) F" and g: "eventually (\x. norm(g x) \ B) F" shows "((\z. f z * g z) \ 0) F" proof - have "((\x. norm (f x) * norm (g x)) \ 0) F" proof (rule Lim_null_comparison) show "\\<^sub>F x in F. norm (norm (f x) * norm (g x)) \ norm (f x) * B" by (simp add: eventually_mono [OF g] mult_left_mono) show "((\x. norm (f x) * B) \ 0) F" by (simp add: f tendsto_mult_left_zero tendsto_norm_zero) qed then show ?thesis by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult) qed lemma lim_null_mult_left_bounded: fixes f :: "'a \ 'b::real_normed_div_algebra" assumes g: "eventually (\x. norm(g x) \ B) F" and f: "(f \ 0) F" shows "((\z. g z * f z) \ 0) F" proof - have "((\x. norm (g x) * norm (f x)) \ 0) F" proof (rule Lim_null_comparison) show "\\<^sub>F x in F. norm (norm (g x) * norm (f x)) \ B * norm (f x)" by (simp add: eventually_mono [OF g] mult_right_mono) show "((\x. B * norm (f x)) \ 0) F" by (simp add: f tendsto_mult_right_zero tendsto_norm_zero) qed then show ?thesis by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult) qed lemma lim_null_scaleR_bounded: assumes f: "(f \ 0) net" and gB: "eventually (\a. f a = 0 \ norm(g a) \ B) net" shows "((\n. f n *\<^sub>R g n) \ 0) net" proof fix \::real assume "0 < \" then have B: "0 < \ / (abs B + 1)" by simp have *: "\f x\ * norm (g x) < \" if f: "\f x\ * (\B\ + 1) < \" and g: "norm (g x) \ B" for x proof - have "\f x\ * norm (g x) \ \f x\ * B" by (simp add: mult_left_mono g) also have "\ \ \f x\ * (\B\ + 1)" by (simp add: mult_left_mono) also have "\ < \" by (rule f) finally show ?thesis . qed have "\x. \\f x\ < \ / (\B\ + 1); norm (g x) \ B\ \ \f x\ * norm (g x) < \" by (simp add: "*" pos_less_divide_eq) then show "\\<^sub>F x in net. dist (f x *\<^sub>R g x) 0 < \" using \0 < \\ by (auto intro: eventually_mono [OF eventually_conj [OF tendstoD [OF f B] gB]]) qed lemma Lim_norm_ubound: fixes f :: "'a \ 'b::real_normed_vector" assumes "\(trivial_limit net)" "(f \ l) net" "eventually (\x. norm(f x) \ e) net" shows "norm(l) \ e" using assms by (fast intro: tendsto_le tendsto_intros) lemma Lim_norm_lbound: fixes f :: "'a \ 'b::real_normed_vector" assumes "\ trivial_limit net" and "(f \ l) net" and "eventually (\x. e \ norm (f x)) net" shows "e \ norm l" using assms by (fast intro: tendsto_le tendsto_intros) text\Limit under bilinear function\ lemma Lim_bilinear: assumes "(f \ l) net" and "(g \ m) net" and "bounded_bilinear h" shows "((\x. h (f x) (g x)) \ (h l m)) net" using \bounded_bilinear h\ \(f \ l) net\ \(g \ m) net\ by (rule bounded_bilinear.tendsto) lemma Lim_at_zero: fixes a :: "'a::real_normed_vector" and l :: "'b::topological_space" shows "(f \ l) (at a) \ ((\x. f(a + x)) \ l) (at 0)" using LIM_offset_zero LIM_offset_zero_cancel .. subsection\<^marker>\tag unimportant\ \Limit Point of Filter\ lemma netlimit_at_vector: fixes a :: "'a::real_normed_vector" shows "netlimit (at a) = a" proof (cases "\x. x \ a") case True then obtain x where x: "x \ a" .. have "\d. 0 < d \ \x. x \ a \ norm (x - a) < d" by (rule_tac x="a + scaleR (d / 2) (sgn (x - a))" in exI) (simp add: norm_sgn sgn_zero_iff x) then have "\ trivial_limit (at a)" by (auto simp: trivial_limit_def eventually_at dist_norm) then show ?thesis by (rule Lim_ident_at [of a UNIV]) qed simp subsection \Boundedness\ lemma continuous_on_closure_norm_le: fixes f :: "'a::metric_space \ 'b::real_normed_vector" assumes "continuous_on (closure s) f" and "\y \ s. norm(f y) \ b" and "x \ (closure s)" shows "norm (f x) \ b" proof - have *: "f ` s \ cball 0 b" using assms(2)[unfolded mem_cball_0[symmetric]] by auto show ?thesis by (meson "*" assms(1) assms(3) closed_cball image_closure_subset image_subset_iff mem_cball_0) qed lemma bounded_pos: "bounded S \ (\b>0. \x\ S. norm x \ b)" unfolding bounded_iff by (meson less_imp_le not_le order_trans zero_less_one) lemma bounded_pos_less: "bounded S \ (\b>0. \x\ S. norm x < b)" by (metis bounded_pos le_less_trans less_imp_le linordered_field_no_ub) lemma Bseq_eq_bounded: fixes f :: "nat \ 'a::real_normed_vector" shows "Bseq f \ bounded (range f)" unfolding Bseq_def bounded_pos by auto lemma bounded_linear_image: assumes "bounded S" and "bounded_linear f" shows "bounded (f ` S)" proof - from assms(1) obtain b where "b > 0" and b: "\x\S. norm x \ b" unfolding bounded_pos by auto from assms(2) obtain B where B: "B > 0" "\x. norm (f x) \ B * norm x" using bounded_linear.pos_bounded by (auto simp: ac_simps) show ?thesis unfolding bounded_pos proof (intro exI, safe) show "norm (f x) \ B * b" if "x \ S" for x by (meson B b less_imp_le mult_left_mono order_trans that) qed (use \b > 0\ \B > 0\ in auto) qed lemma bounded_scaling: fixes S :: "'a::real_normed_vector set" shows "bounded S \ bounded ((\x. c *\<^sub>R x) ` S)" by (simp add: bounded_linear_image bounded_linear_scaleR_right) lemma bounded_scaleR_comp: fixes f :: "'a \ 'b::real_normed_vector" assumes "bounded (f ` S)" shows "bounded ((\x. r *\<^sub>R f x) ` S)" using bounded_scaling[of "f ` S" r] assms by (auto simp: image_image) lemma bounded_translation: fixes S :: "'a::real_normed_vector set" assumes "bounded S" shows "bounded ((\x. a + x) ` S)" proof - from assms obtain b where b: "b > 0" "\x\S. norm x \ b" unfolding bounded_pos by auto { fix x assume "x \ S" then have "norm (a + x) \ b + norm a" using norm_triangle_ineq[of a x] b by auto } then show ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) and add_strict_increasing[of b 0 "norm a"] by (auto intro!: exI[of _ "b + norm a"]) qed lemma bounded_translation_minus: fixes S :: "'a::real_normed_vector set" shows "bounded S \ bounded ((\x. x - a) ` S)" using bounded_translation [of S "-a"] by simp lemma bounded_uminus [simp]: fixes X :: "'a::real_normed_vector set" shows "bounded (uminus ` X) \ bounded X" by (auto simp: bounded_def dist_norm; rule_tac x="-x" in exI; force simp: add.commute norm_minus_commute) lemma uminus_bounded_comp [simp]: fixes f :: "'a \ 'b::real_normed_vector" shows "bounded ((\x. - f x) ` S) \ bounded (f ` S)" using bounded_uminus[of "f ` S"] by (auto simp: image_image) lemma bounded_plus_comp: fixes f g::"'a \ 'b::real_normed_vector" assumes "bounded (f ` S)" assumes "bounded (g ` S)" shows "bounded ((\x. f x + g x) ` S)" proof - { fix B C assume "\x. x\S \ norm (f x) \ B" "\x. x\S \ norm (g x) \ C" then have "\x. x \ S \ norm (f x + g x) \ B + C" by (auto intro!: norm_triangle_le add_mono) } then show ?thesis using assms by (fastforce simp: bounded_iff) qed lemma bounded_plus: fixes S ::"'a::real_normed_vector set" assumes "bounded S" "bounded T" shows "bounded ((\(x,y). x + y) ` (S \ T))" using bounded_plus_comp [of fst "S \ T" snd] assms by (auto simp: split_def split: if_split_asm) lemma bounded_minus_comp: "bounded (f ` S) \ bounded (g ` S) \ bounded ((\x. f x - g x) ` S)" for f g::"'a \ 'b::real_normed_vector" using bounded_plus_comp[of "f" S "\x. - g x"] by auto lemma bounded_minus: fixes S ::"'a::real_normed_vector set" assumes "bounded S" "bounded T" shows "bounded ((\(x,y). x - y) ` (S \ T))" using bounded_minus_comp [of fst "S \ T" snd] assms by (auto simp: split_def split: if_split_asm) +lemma bounded_sums: + fixes S :: "'a::real_normed_vector set" + assumes "bounded S" and "bounded T" + shows "bounded (\x\ S. \y \ T. {x + y})" + using assms by (simp add: bounded_iff) (meson norm_triangle_mono) + +lemma bounded_differences: + fixes S :: "'a::real_normed_vector set" + assumes "bounded S" and "bounded T" + shows "bounded (\x\ S. \y \ T. {x - y})" + using assms by (simp add: bounded_iff) (meson add_mono norm_triangle_le_diff) + lemma not_bounded_UNIV[simp]: "\ bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)" proof (auto simp: bounded_pos not_le) obtain x :: 'a where "x \ 0" using perfect_choose_dist [OF zero_less_one] by fast fix b :: real assume b: "b >0" have b1: "b +1 \ 0" using b by simp with \x \ 0\ have "b < norm (scaleR (b + 1) (sgn x))" by (simp add: norm_sgn) then show "\x::'a. b < norm x" .. qed corollary cobounded_imp_unbounded: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded (- S) \ \ bounded S" using bounded_Un [of S "-S"] by (simp) subsection\<^marker>\tag unimportant\\Relations among convergence and absolute convergence for power series\ lemma summable_imp_bounded: fixes f :: "nat \ 'a::real_normed_vector" shows "summable f \ bounded (range f)" by (frule summable_LIMSEQ_zero) (simp add: convergent_imp_bounded) lemma summable_imp_sums_bounded: "summable f \ bounded (range (\n. sum f {.. 'a::{real_normed_div_algebra,banach}" and w :: 'a assumes sum: "summable (\n. a n * z ^ n)" and no: "norm w < norm z" shows "summable (\n. of_real(norm(a n)) * w ^ n)" proof - obtain M where M: "\x. norm (a x * z ^ x) \ M" using summable_imp_bounded [OF sum] by (force simp: bounded_iff) show ?thesis proof (rule series_comparison_complex) have "\n. norm (a n) * norm z ^ n \ M" by (metis (no_types) M norm_mult norm_power) then show "summable (\n. complex_of_real (norm (a n) * norm w ^ n))" using Abel_lemma no norm_ge_zero summable_of_real by blast qed (auto simp: norm_mult norm_power) qed subsection \Normed spaces with the Heine-Borel property\ lemma not_compact_UNIV[simp]: fixes s :: "'a::{real_normed_vector,perfect_space,heine_borel} set" shows "\ compact (UNIV::'a set)" by (simp add: compact_eq_bounded_closed) lemma not_compact_space_euclideanreal [simp]: "\ compact_space euclideanreal" by (simp add: compact_space_def) text\Representing sets as the union of a chain of compact sets.\ lemma closed_Union_compact_subsets: fixes S :: "'a::{heine_borel,real_normed_vector} set" assumes "closed S" obtains F where "\n. compact(F n)" "\n. F n \ S" "\n. F n \ F(Suc n)" "(\n. F n) = S" "\K. \compact K; K \ S\ \ \N. \n \ N. K \ F n" proof show "compact (S \ cball 0 (of_nat n))" for n using assms compact_eq_bounded_closed by auto next show "(\n. S \ cball 0 (real n)) = S" by (auto simp: real_arch_simple) next fix K :: "'a set" assume "compact K" "K \ S" then obtain N where "K \ cball 0 N" by (meson bounded_pos mem_cball_0 compact_imp_bounded subsetI) then show "\N. \n\N. K \ S \ cball 0 (real n)" by (metis of_nat_le_iff Int_subset_iff \K \ S\ real_arch_simple subset_cball subset_trans) qed auto subsection \Intersecting chains of compact sets and the Baire property\ proposition bounded_closed_chain: fixes \ :: "'a::heine_borel set set" assumes "B \ \" "bounded B" and \: "\S. S \ \ \ closed S" and "{} \ \" and chain: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "\\ \ {}" proof - have "B \ \\ \ {}" proof (rule compact_imp_fip) show "compact B" "\T. T \ \ \ closed T" by (simp_all add: assms compact_eq_bounded_closed) show "\finite \; \ \ \\ \ B \ \\ \ {}" for \ proof (induction \ rule: finite_induct) case empty with assms show ?case by force next case (insert U \) then have "U \ \" and ne: "B \ \\ \ {}" by auto then consider "B \ U" | "U \ B" using \B \ \\ chain by blast then show ?case proof cases case 1 then show ?thesis using Int_left_commute ne by auto next case 2 have "U \ {}" using \U \ \\ \{} \ \\ by blast moreover have False if "\x. x \ U \ \Y\\. x \ Y" proof - have "\x. x \ U \ \Y\\. Y \ U" by (metis chain contra_subsetD insert.prems insert_subset that) then obtain Y where "Y \ \" "Y \ U" by (metis all_not_in_conv \U \ {}\) moreover obtain x where "x \ \\" by (metis Int_emptyI ne) ultimately show ?thesis by (metis Inf_lower subset_eq that) qed with 2 show ?thesis by blast qed qed qed then show ?thesis by blast qed corollary compact_chain: fixes \ :: "'a::heine_borel set set" assumes "\S. S \ \ \ compact S" "{} \ \" "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "\ \ \ {}" proof (cases "\ = {}") case True then show ?thesis by auto next case False show ?thesis by (metis False all_not_in_conv assms compact_imp_bounded compact_imp_closed bounded_closed_chain) qed lemma compact_nest: fixes F :: "'a::linorder \ 'b::heine_borel set" assumes F: "\n. compact(F n)" "\n. F n \ {}" and mono: "\m n. m \ n \ F n \ F m" shows "\(range F) \ {}" proof - have *: "\S T. S \ range F \ T \ range F \ S \ T \ T \ S" by (metis mono image_iff le_cases) show ?thesis using F by (intro compact_chain [OF _ _ *]; blast dest: *) qed text\The Baire property of dense sets\ theorem Baire: fixes S::"'a::{real_normed_vector,heine_borel} set" assumes "closed S" "countable \" and ope: "\T. T \ \ \ openin (top_of_set S) T \ S \ closure T" shows "S \ closure(\\)" proof (cases "\ = {}") case True then show ?thesis using closure_subset by auto next let ?g = "from_nat_into \" case False then have gin: "?g n \ \" for n by (simp add: from_nat_into) show ?thesis proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" obtain TF where opeF: "\n. openin (top_of_set S) (TF n)" and ne: "\n. TF n \ {}" and subg: "\n. S \ closure(TF n) \ ?g n" and subball: "\n. closure(TF n) \ ball x e" and decr: "\n. TF(Suc n) \ TF n" proof - have *: "\Y. (openin (top_of_set S) Y \ Y \ {} \ S \ closure Y \ ?g n \ closure Y \ ball x e) \ Y \ U" if opeU: "openin (top_of_set S) U" and "U \ {}" and cloU: "closure U \ ball x e" for U n proof - obtain T where T: "open T" "U = T \ S" using \openin (top_of_set S) U\ by (auto simp: openin_subtopology) with \U \ {}\ have "T \ closure (?g n) \ {}" using gin ope by fastforce then have "T \ ?g n \ {}" using \open T\ open_Int_closure_eq_empty by blast then obtain y where "y \ U" "y \ ?g n" using T ope [of "?g n", OF gin] by (blast dest: openin_imp_subset) moreover have "openin (top_of_set S) (U \ ?g n)" using gin ope opeU by blast ultimately obtain d where U: "U \ ?g n \ S" and "d > 0" and d: "ball y d \ S \ U \ ?g n" by (force simp: openin_contains_ball) show ?thesis proof (intro exI conjI) show "openin (top_of_set S) (S \ ball y (d/2))" by (simp add: openin_open_Int) show "S \ ball y (d/2) \ {}" using \0 < d\ \y \ U\ opeU openin_imp_subset by fastforce have "S \ closure (S \ ball y (d/2)) \ S \ closure (ball y (d/2))" using closure_mono by blast also have "... \ ?g n" using \d > 0\ d by force finally show "S \ closure (S \ ball y (d/2)) \ ?g n" . have "closure (S \ ball y (d/2)) \ S \ ball y d" proof - have "closure (ball y (d/2)) \ ball y d" using \d > 0\ by auto then have "closure (S \ ball y (d/2)) \ ball y d" by (meson closure_mono inf.cobounded2 subset_trans) then show ?thesis by (simp add: \closed S\ closure_minimal) qed also have "... \ ball x e" using cloU closure_subset d by blast finally show "closure (S \ ball y (d/2)) \ ball x e" . show "S \ ball y (d/2) \ U" using ball_divide_subset_numeral d by blast qed qed let ?\ = "\n X. openin (top_of_set S) X \ X \ {} \ S \ closure X \ ?g n \ closure X \ ball x e" have "closure (S \ ball x (e/2)) \ closure(ball x (e/2))" by (simp add: closure_mono) also have "... \ ball x e" using \e > 0\ by auto finally have "closure (S \ ball x (e/2)) \ ball x e" . moreover have"openin (top_of_set S) (S \ ball x (e/2))" "S \ ball x (e/2) \ {}" using \0 < e\ \x \ S\ by auto ultimately obtain Y where Y: "?\ 0 Y \ Y \ S \ ball x (e/2)" using * [of "S \ ball x (e/2)" 0] by metis show thesis proof (rule exE [OF dependent_nat_choice]) show "\x. ?\ 0 x" using Y by auto show "\Y. ?\ (Suc n) Y \ Y \ X" if "?\ n X" for X n using that by (blast intro: *) qed (use that in metis) qed have "(\n. S \ closure (TF n)) \ {}" proof (rule compact_nest) show "\n. compact (S \ closure (TF n))" by (metis closed_closure subball bounded_subset_ballI compact_eq_bounded_closed closed_Int_compact [OF \closed S\]) show "\n. S \ closure (TF n) \ {}" by (metis Int_absorb1 opeF \closed S\ closure_eq_empty closure_minimal ne openin_imp_subset) show "\m n. m \ n \ S \ closure (TF n) \ S \ closure (TF m)" by (meson closure_mono decr dual_order.refl inf_mono lift_Suc_antimono_le) qed moreover have "(\n. S \ closure (TF n)) \ {y \ \\. dist y x < e}" proof (clarsimp, intro conjI) fix y assume "y \ S" and y: "\n. y \ closure (TF n)" then show "\T\\. y \ T" by (metis Int_iff from_nat_into_surj [OF \countable \\] subsetD subg) show "dist y x < e" by (metis y dist_commute mem_ball subball subsetCE) qed ultimately show "\y \ \\. dist y x < e" by auto qed qed subsection \Continuity\ subsubsection\<^marker>\tag unimportant\ \Structural rules for uniform continuity\ lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]: fixes g :: "_::metric_space \ _" assumes "uniformly_continuous_on s g" shows "uniformly_continuous_on s (\x. f (g x))" using assms unfolding uniformly_continuous_on_sequentially unfolding dist_norm tendsto_norm_zero_iff diff[symmetric] by (auto intro: tendsto_zero) lemma uniformly_continuous_on_dist[continuous_intros]: fixes f g :: "'a::metric_space \ 'b::metric_space" assumes "uniformly_continuous_on s f" and "uniformly_continuous_on s g" shows "uniformly_continuous_on s (\x. dist (f x) (g x))" proof - { fix a b c d :: 'b have "\dist a b - dist c d\ \ dist a c + dist b d" using dist_triangle2 [of a b c] dist_triangle2 [of b c d] using dist_triangle3 [of c d a] dist_triangle [of a d b] by arith } note le = this { fix x y assume f: "(\n. dist (f (x n)) (f (y n))) \ 0" assume g: "(\n. dist (g (x n)) (g (y n))) \ 0" have "(\n. \dist (f (x n)) (g (x n)) - dist (f (y n)) (g (y n))\) \ 0" by (rule Lim_transform_bound [OF _ tendsto_add_zero [OF f g]], simp add: le) } then show ?thesis using assms unfolding uniformly_continuous_on_sequentially unfolding dist_real_def by simp qed lemma uniformly_continuous_on_cmul_right [continuous_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "uniformly_continuous_on s f \ uniformly_continuous_on s (\x. f x * c)" using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] . lemma uniformly_continuous_on_cmul_left[continuous_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_algebra" assumes "uniformly_continuous_on s f" shows "uniformly_continuous_on s (\x. c * f x)" by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right) lemma uniformly_continuous_on_norm[continuous_intros]: fixes f :: "'a :: metric_space \ 'b :: real_normed_vector" assumes "uniformly_continuous_on s f" shows "uniformly_continuous_on s (\x. norm (f x))" unfolding norm_conv_dist using assms by (intro uniformly_continuous_on_dist uniformly_continuous_on_const) lemma uniformly_continuous_on_cmul[continuous_intros]: fixes f :: "'a::metric_space \ 'b::real_normed_vector" assumes "uniformly_continuous_on s f" shows "uniformly_continuous_on s (\x. c *\<^sub>R f(x))" using bounded_linear_scaleR_right assms by (rule bounded_linear.uniformly_continuous_on) lemma dist_minus: fixes x y :: "'a::real_normed_vector" shows "dist (- x) (- y) = dist x y" unfolding dist_norm minus_diff_minus norm_minus_cancel .. lemma uniformly_continuous_on_minus[continuous_intros]: fixes f :: "'a::metric_space \ 'b::real_normed_vector" shows "uniformly_continuous_on s f \ uniformly_continuous_on s (\x. - f x)" unfolding uniformly_continuous_on_def dist_minus . lemma uniformly_continuous_on_add[continuous_intros]: fixes f g :: "'a::metric_space \ 'b::real_normed_vector" assumes "uniformly_continuous_on s f" and "uniformly_continuous_on s g" shows "uniformly_continuous_on s (\x. f x + g x)" using assms unfolding uniformly_continuous_on_sequentially unfolding dist_norm tendsto_norm_zero_iff add_diff_add by (auto intro: tendsto_add_zero) lemma uniformly_continuous_on_diff[continuous_intros]: fixes f :: "'a::metric_space \ 'b::real_normed_vector" assumes "uniformly_continuous_on s f" and "uniformly_continuous_on s g" shows "uniformly_continuous_on s (\x. f x - g x)" using assms uniformly_continuous_on_add [of s f "- g"] by (simp add: fun_Compl_def uniformly_continuous_on_minus) subsection\<^marker>\tag unimportant\ \Arithmetic Preserves Topological Properties\ lemma open_scaling[intro]: fixes s :: "'a::real_normed_vector set" assumes "c \ 0" and "open s" shows "open((\x. c *\<^sub>R x) ` s)" proof - { fix x assume "x \ s" then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] by auto have "e * \c\ > 0" using assms(1)[unfolded zero_less_abs_iff[symmetric]] \e>0\ by auto moreover { fix y assume "dist y (c *\<^sub>R x) < e * \c\" then have "norm (c *\<^sub>R ((1 / c) *\<^sub>R y - x)) < e * norm c" by (simp add: \c \ 0\ dist_norm scale_right_diff_distrib) then have "norm ((1 / c) *\<^sub>R y - x) < e" by (simp add: \c \ 0\) then have "y \ (*\<^sub>R) c ` s" using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "(*\<^sub>R) c"] by (simp add: \c \ 0\ dist_norm e) } ultimately have "\e>0. \x'. dist x' (c *\<^sub>R x) < e \ x' \ (*\<^sub>R) c ` s" by (rule_tac x="e * \c\" in exI, auto) } then show ?thesis unfolding open_dist by auto qed lemma minus_image_eq_vimage: fixes A :: "'a::ab_group_add set" shows "(\x. - x) ` A = (\x. - x) -` A" by (auto intro!: image_eqI [where f="\x. - x"]) lemma open_negations: fixes S :: "'a::real_normed_vector set" shows "open S \ open ((\x. - x) ` S)" using open_scaling [of "- 1" S] by simp lemma open_translation: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open((\x. a + x) ` S)" proof - { fix x have "continuous (at x) (\x. x - a)" by (intro continuous_diff continuous_ident continuous_const) } moreover have "{x. x - a \ S} = (+) a ` S" by force ultimately show ?thesis by (metis assms continuous_open_vimage vimage_def) qed lemma open_translation_subtract: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open ((\x. x - a) ` S)" using assms open_translation [of S "- a"] by (simp cong: image_cong_simp) lemma open_neg_translation: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open((\x. a - x) ` S)" using open_translation[OF open_negations[OF assms], of a] by (auto simp: image_image) lemma open_affinity: fixes S :: "'a::real_normed_vector set" assumes "open S" "c \ 0" shows "open ((\x. a + c *\<^sub>R x) ` S)" proof - have *: "(\x. a + c *\<^sub>R x) = (\x. a + x) \ (\x. c *\<^sub>R x)" unfolding o_def .. have "(+) a ` (*\<^sub>R) c ` S = ((+) a \ (*\<^sub>R) c) ` S" by auto then show ?thesis using assms open_translation[of "(*\<^sub>R) c ` S" a] unfolding * by auto qed lemma interior_translation: "interior ((+) a ` S) = (+) a ` (interior S)" for S :: "'a::real_normed_vector set" proof (rule set_eqI, rule) fix x assume "x \ interior ((+) a ` S)" then obtain e where "e > 0" and e: "ball x e \ (+) a ` S" unfolding mem_interior by auto then have "ball (x - a) e \ S" unfolding subset_eq Ball_def mem_ball dist_norm by (auto simp: diff_diff_eq) then show "x \ (+) a ` interior S" unfolding image_iff by (metis \0 < e\ add.commute diff_add_cancel mem_interior) next fix x assume "x \ (+) a ` interior S" then obtain y e where "e > 0" and e: "ball y e \ S" and y: "x = a + y" unfolding image_iff Bex_def mem_interior by auto { fix z have *: "a + y - z = y + a - z" by auto assume "z \ ball x e" then have "z - a \ S" using e[unfolded subset_eq, THEN bspec[where x="z - a"]] unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 * by auto then have "z \ (+) a ` S" unfolding image_iff by (auto intro!: bexI[where x="z - a"]) } then have "ball x e \ (+) a ` S" unfolding subset_eq by auto then show "x \ interior ((+) a ` S)" unfolding mem_interior using \e > 0\ by auto qed lemma interior_translation_subtract: "interior ((\x. x - a) ` S) = (\x. x - a) ` interior S" for S :: "'a::real_normed_vector set" using interior_translation [of "- a"] by (simp cong: image_cong_simp) lemma compact_scaling: fixes s :: "'a::real_normed_vector set" assumes "compact s" shows "compact ((\x. c *\<^sub>R x) ` s)" proof - let ?f = "\x. scaleR c x" have *: "bounded_linear ?f" by (rule bounded_linear_scaleR_right) show ?thesis using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] using linear_continuous_at[OF *] assms by auto qed lemma compact_negations: fixes s :: "'a::real_normed_vector set" assumes "compact s" shows "compact ((\x. - x) ` s)" using compact_scaling [OF assms, of "- 1"] by auto lemma compact_sums: fixes s t :: "'a::real_normed_vector set" assumes "compact s" and "compact t" shows "compact {x + y | x y. x \ s \ y \ t}" proof - have *: "{x + y | x y. x \ s \ y \ t} = (\z. fst z + snd z) ` (s \ t)" by (fastforce simp: image_iff) have "continuous_on (s \ t) (\z. fst z + snd z)" unfolding continuous_on by (rule ballI) (intro tendsto_intros) then show ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto qed lemma compact_differences: fixes s t :: "'a::real_normed_vector set" assumes "compact s" and "compact t" shows "compact {x - y | x y. x \ s \ y \ t}" proof- have "{x - y | x y. x\s \ y \ t} = {x + y | x y. x \ s \ y \ (uminus ` t)}" using diff_conv_add_uminus by force then show ?thesis using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto qed +lemma compact_sums': + fixes S :: "'a::real_normed_vector set" + assumes "compact S" and "compact T" + shows "compact (\x\ S. \y \ T. {x + y})" +proof - + have "(\x\S. \y\T. {x + y}) = {x + y |x y. x \ S \ y \ T}" + by blast + then show ?thesis + using compact_sums [OF assms] by simp +qed + +lemma compact_differences': + fixes S :: "'a::real_normed_vector set" + assumes "compact S" and "compact T" + shows "compact (\x\ S. \y \ T. {x - y})" +proof - + have "(\x\S. \y\T. {x - y}) = {x - y |x y. x \ S \ y \ T}" + by blast + then show ?thesis + using compact_differences [OF assms] by simp +qed + lemma compact_translation: "compact ((+) a ` s)" if "compact s" for s :: "'a::real_normed_vector set" proof - have "{x + y |x y. x \ s \ y \ {a}} = (\x. a + x) ` s" by auto then show ?thesis using compact_sums [OF that compact_sing [of a]] by auto qed lemma compact_translation_subtract: "compact ((\x. x - a) ` s)" if "compact s" for s :: "'a::real_normed_vector set" using that compact_translation [of s "- a"] by (simp cong: image_cong_simp) lemma compact_affinity: fixes s :: "'a::real_normed_vector set" assumes "compact s" shows "compact ((\x. a + c *\<^sub>R x) ` s)" proof - have "(+) a ` (*\<^sub>R) c ` s = (\x. a + c *\<^sub>R x) ` s" by auto then show ?thesis using compact_translation[OF compact_scaling[OF assms], of a c] by auto qed lemma closed_scaling: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "closed ((\x. c *\<^sub>R x) ` S)" proof (cases "c = 0") case True then show ?thesis by (auto simp: image_constant_conv) next case False from assms have "closed ((\x. inverse c *\<^sub>R x) -` S)" by (simp add: continuous_closed_vimage) also have "(\x. inverse c *\<^sub>R x) -` S = (\x. c *\<^sub>R x) ` S" using \c \ 0\ by (auto elim: image_eqI [rotated]) finally show ?thesis . qed lemma closed_negations: fixes S :: "'a::real_normed_vector set" assumes "closed S" shows "closed ((\x. -x) ` S)" using closed_scaling[OF assms, of "- 1"] by simp lemma compact_closed_sums: fixes S :: "'a::real_normed_vector set" assumes "compact S" and "closed T" shows "closed (\x\ S. \y \ T. {x + y})" proof - let ?S = "{x + y |x y. x \ S \ y \ T}" { fix x l assume as: "\n. x n \ ?S" "(x \ l) sequentially" from as(1) obtain f where f: "\n. x n = fst (f n) + snd (f n)" "\n. fst (f n) \ S" "\n. snd (f n) \ T" using choice[of "\n y. x n = (fst y) + (snd y) \ fst y \ S \ snd y \ T"] by auto obtain l' r where "l'\S" and r: "strict_mono r" and lr: "(((\n. fst (f n)) \ r) \ l') sequentially" using assms(1)[unfolded compact_def, THEN spec[where x="\ n. fst (f n)"]] using f(2) by auto have "((\n. snd (f (r n))) \ l - l') sequentially" using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1) unfolding o_def by auto then have "l - l' \ T" using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\ n. snd (f (r n))"], THEN spec[where x="l - l'"]] using f(3) by auto then have "l \ ?S" using \l' \ S\ by force } moreover have "?S = (\x\ S. \y \ T. {x + y})" by force ultimately show ?thesis unfolding closed_sequential_limits by (metis (no_types, lifting)) qed lemma closed_compact_sums: fixes S T :: "'a::real_normed_vector set" assumes "closed S" "compact T" shows "closed (\x\ S. \y \ T. {x + y})" proof - have "(\x\ T. \y \ S. {x + y}) = (\x\ S. \y \ T. {x + y})" by auto then show ?thesis using compact_closed_sums[OF assms(2,1)] by simp qed lemma compact_closed_differences: fixes S T :: "'a::real_normed_vector set" assumes "compact S" "closed T" shows "closed (\x\ S. \y \ T. {x - y})" proof - have "(\x\ S. \y \ uminus ` T. {x + y}) = (\x\ S. \y \ T. {x - y})" by force then show ?thesis by (metis assms closed_negations compact_closed_sums) qed lemma closed_compact_differences: fixes S T :: "'a::real_normed_vector set" assumes "closed S" "compact T" shows "closed (\x\ S. \y \ T. {x - y})" proof - have "(\x\ S. \y \ uminus ` T. {x + y}) = {x - y |x y. x \ S \ y \ T}" by auto then show ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp qed lemma closed_translation: "closed ((+) a ` S)" if "closed S" for a :: "'a::real_normed_vector" proof - have "(\x\ {a}. \y \ S. {x + y}) = ((+) a ` S)" by auto then show ?thesis using compact_closed_sums [OF compact_sing [of a] that] by auto qed lemma closed_translation_subtract: "closed ((\x. x - a) ` S)" if "closed S" for a :: "'a::real_normed_vector" using that closed_translation [of S "- a"] by (simp cong: image_cong_simp) lemma closure_translation: "closure ((+) a ` s) = (+) a ` closure s" for a :: "'a::real_normed_vector" proof - have *: "(+) a ` (- s) = - (+) a ` s" by (auto intro!: image_eqI [where x = "x - a" for x]) show ?thesis using interior_translation [of a "- s", symmetric] by (simp add: closure_interior translation_Compl *) qed lemma closure_translation_subtract: "closure ((\x. x - a) ` s) = (\x. x - a) ` closure s" for a :: "'a::real_normed_vector" using closure_translation [of "- a" s] by (simp cong: image_cong_simp) lemma frontier_translation: "frontier ((+) a ` s) = (+) a ` frontier s" for a :: "'a::real_normed_vector" by (auto simp add: frontier_def translation_diff interior_translation closure_translation) lemma frontier_translation_subtract: "frontier ((+) a ` s) = (+) a ` frontier s" for a :: "'a::real_normed_vector" by (auto simp add: frontier_def translation_diff interior_translation closure_translation) lemma sphere_translation: "sphere (a + c) r = (+) a ` sphere c r" for a :: "'n::real_normed_vector" by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x]) lemma sphere_translation_subtract: "sphere (c - a) r = (\x. x - a) ` sphere c r" for a :: "'n::real_normed_vector" using sphere_translation [of "- a" c] by (simp cong: image_cong_simp) lemma cball_translation: "cball (a + c) r = (+) a ` cball c r" for a :: "'n::real_normed_vector" by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x]) lemma cball_translation_subtract: "cball (c - a) r = (\x. x - a) ` cball c r" for a :: "'n::real_normed_vector" using cball_translation [of "- a" c] by (simp cong: image_cong_simp) lemma ball_translation: "ball (a + c) r = (+) a ` ball c r" for a :: "'n::real_normed_vector" by (auto simp: dist_norm algebra_simps intro!: image_eqI [where x = "x - a" for x]) lemma ball_translation_subtract: "ball (c - a) r = (\x. x - a) ` ball c r" for a :: "'n::real_normed_vector" using ball_translation [of "- a" c] by (simp cong: image_cong_simp) subsection\<^marker>\tag unimportant\\Homeomorphisms\ lemma homeomorphic_scaling: fixes S :: "'a::real_normed_vector set" assumes "c \ 0" shows "S homeomorphic ((\x. c *\<^sub>R x) ` S)" unfolding homeomorphic_minimal apply (rule_tac x="\x. c *\<^sub>R x" in exI) apply (rule_tac x="\x. (1 / c) *\<^sub>R x" in exI) using assms by (auto simp: continuous_intros) lemma homeomorphic_translation: fixes S :: "'a::real_normed_vector set" shows "S homeomorphic ((\x. a + x) ` S)" unfolding homeomorphic_minimal apply (rule_tac x="\x. a + x" in exI) apply (rule_tac x="\x. -a + x" in exI) by (auto simp: continuous_intros) lemma homeomorphic_affinity: fixes S :: "'a::real_normed_vector set" assumes "c \ 0" shows "S homeomorphic ((\x. a + c *\<^sub>R x) ` S)" proof - have *: "(+) a ` (*\<^sub>R) c ` S = (\x. a + c *\<^sub>R x) ` S" by auto show ?thesis by (metis "*" assms homeomorphic_scaling homeomorphic_trans homeomorphic_translation) qed lemma homeomorphic_balls: fixes a b ::"'a::real_normed_vector" assumes "0 < d" "0 < e" shows "(ball a d) homeomorphic (ball b e)" (is ?th) and "(cball a d) homeomorphic (cball b e)" (is ?cth) proof - show ?th unfolding homeomorphic_minimal apply(rule_tac x="\x. b + (e/d) *\<^sub>R (x - a)" in exI) apply(rule_tac x="\x. a + (d/e) *\<^sub>R (x - b)" in exI) using assms by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_less_eq) show ?cth unfolding homeomorphic_minimal apply(rule_tac x="\x. b + (e/d) *\<^sub>R (x - a)" in exI) apply(rule_tac x="\x. a + (d/e) *\<^sub>R (x - b)" in exI) using assms by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_le_eq) qed lemma homeomorphic_spheres: fixes a b ::"'a::real_normed_vector" assumes "0 < d" "0 < e" shows "(sphere a d) homeomorphic (sphere b e)" unfolding homeomorphic_minimal apply(rule_tac x="\x. b + (e/d) *\<^sub>R (x - a)" in exI) apply(rule_tac x="\x. a + (d/e) *\<^sub>R (x - b)" in exI) using assms by (auto intro!: continuous_intros simp: dist_commute dist_norm pos_divide_less_eq) lemma homeomorphic_ball01_UNIV: "ball (0::'a::real_normed_vector) 1 homeomorphic (UNIV:: 'a set)" (is "?B homeomorphic ?U") proof have "x \ (\z. z /\<^sub>R (1 - norm z)) ` ball 0 1" for x::'a apply (rule_tac x="x /\<^sub>R (1 + norm x)" in image_eqI) apply (auto simp: field_split_simps) using norm_ge_zero [of x] apply linarith+ done then show "(\z::'a. z /\<^sub>R (1 - norm z)) ` ?B = ?U" by blast have "x \ range (\z. (1 / (1 + norm z)) *\<^sub>R z)" if "norm x < 1" for x::'a using that by (rule_tac x="x /\<^sub>R (1 - norm x)" in image_eqI) (auto simp: field_split_simps) then show "(\z::'a. z /\<^sub>R (1 + norm z)) ` ?U = ?B" by (force simp: field_split_simps dest: add_less_zeroD) show "continuous_on (ball 0 1) (\z. z /\<^sub>R (1 - norm z))" by (rule continuous_intros | force)+ have 0: "\z. 1 + norm z \ 0" by (metis (no_types) le_add_same_cancel1 norm_ge_zero not_one_le_zero) then show "continuous_on UNIV (\z. z /\<^sub>R (1 + norm z))" by (auto intro!: continuous_intros) show "\x. x \ ball 0 1 \ x /\<^sub>R (1 - norm x) /\<^sub>R (1 + norm (x /\<^sub>R (1 - norm x))) = x" by (auto simp: field_split_simps) show "\y. y /\<^sub>R (1 + norm y) /\<^sub>R (1 - norm (y /\<^sub>R (1 + norm y))) = y" using 0 by (auto simp: field_split_simps) qed proposition homeomorphic_ball_UNIV: fixes a ::"'a::real_normed_vector" assumes "0 < r" shows "ball a r homeomorphic (UNIV:: 'a set)" using assms homeomorphic_ball01_UNIV homeomorphic_balls(1) homeomorphic_trans zero_less_one by blast subsection\<^marker>\tag unimportant\ \Discrete\ lemma finite_implies_discrete: fixes S :: "'a::topological_space set" assumes "finite (f ` S)" shows "(\x \ S. \e>0. \y. y \ S \ f y \ f x \ e \ norm (f y - f x))" proof - have "\e>0. \y. y \ S \ f y \ f x \ e \ norm (f y - f x)" if "x \ S" for x proof (cases "f ` S - {f x} = {}") case True with zero_less_numeral show ?thesis by (fastforce simp add: Set.image_subset_iff cong: conj_cong) next case False then obtain z where "z \ S" "f z \ f x" by blast moreover have finn: "finite {norm (z - f x) |z. z \ f ` S - {f x}}" using assms by simp ultimately have *: "0 < Inf{norm(z - f x) | z. z \ f ` S - {f x}}" by (force intro: finite_imp_less_Inf) show ?thesis by (force intro!: * cInf_le_finite [OF finn]) qed with assms show ?thesis by blast qed subsection\<^marker>\tag unimportant\ \Completeness of "Isometry" (up to constant bounds)\ lemma cauchy_isometric:\ \TODO: rename lemma to \Cauchy_isometric\\ assumes e: "e > 0" and s: "subspace s" and f: "bounded_linear f" and normf: "\x\s. norm (f x) \ e * norm x" and xs: "\n. x n \ s" and cf: "Cauchy (f \ x)" shows "Cauchy x" proof - interpret f: bounded_linear f by fact have "\N. \n\N. norm (x n - x N) < d" if "d > 0" for d :: real proof - from that obtain N where N: "\n\N. norm (f (x n) - f (x N)) < e * d" using cf[unfolded Cauchy_def o_def dist_norm, THEN spec[where x="e*d"]] e by auto have "norm (x n - x N) < d" if "n \ N" for n proof - have "e * norm (x n - x N) \ norm (f (x n - x N))" using subspace_diff[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] using normf[THEN bspec[where x="x n - x N"]] by auto also have "norm (f (x n - x N)) < e * d" using \N \ n\ N unfolding f.diff[symmetric] by auto finally show ?thesis using \e>0\ by simp qed then show ?thesis by auto qed then show ?thesis by (simp add: Cauchy_altdef2 dist_norm) qed lemma complete_isometric_image: assumes "0 < e" and s: "subspace s" and f: "bounded_linear f" and normf: "\x\s. norm(f x) \ e * norm(x)" and cs: "complete s" shows "complete (f ` s)" proof - have "\l\f ` s. (g \ l) sequentially" if as:"\n::nat. g n \ f ` s" and cfg:"Cauchy g" for g proof - from that obtain x where "\n. x n \ s \ g n = f (x n)" using choice[of "\ n xa. xa \ s \ g n = f xa"] by auto then have x: "\n. x n \ s" "\n. g n = f (x n)" by auto then have "f \ x = g" by (simp add: fun_eq_iff) then obtain l where "l\s" and l:"(x \ l) sequentially" using cs[unfolded complete_def, THEN spec[where x=x]] using cauchy_isometric[OF \0 < e\ s f normf] and cfg and x(1) by auto then show ?thesis using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] by (auto simp: \f \ x = g\) qed then show ?thesis unfolding complete_def by auto qed subsection \Connected Normed Spaces\ lemma compact_components: fixes s :: "'a::heine_borel set" shows "\compact s; c \ components s\ \ compact c" by (meson bounded_subset closed_components in_components_subset compact_eq_bounded_closed) lemma discrete_subset_disconnected: fixes S :: "'a::topological_space set" fixes t :: "'b::real_normed_vector set" assumes conf: "continuous_on S f" and no: "\x. x \ S \ \e>0. \y. y \ S \ f y \ f x \ e \ norm (f y - f x)" shows "f ` S \ {y. connected_component_set (f ` S) y = {y}}" proof - { fix x assume x: "x \ S" then obtain e where "e>0" and ele: "\y. \y \ S; f y \ f x\ \ e \ norm (f y - f x)" using conf no [OF x] by auto then have e2: "0 \ e/2" by simp define F where "F \ connected_component_set (f ` S) (f x)" have False if "y \ S" and ccs: "f y \ F" and not: "f y \ f x" for y proof - define C where "C \ cball (f x) (e/2)" define D where "D \ - ball (f x) e" have disj: "C \ D = {}" unfolding C_def D_def using \0 < e\ by fastforce moreover have FCD: "F \ C \ D" proof - have "t \ C \ t \ D" if "t \ F" for t proof - obtain y where "y \ S" "t = f y" using F_def \t \ F\ connected_component_in by blast then show ?thesis by (metis C_def ComplI D_def centre_in_cball dist_norm e2 ele mem_ball norm_minus_commute not_le) qed then show ?thesis by auto qed ultimately have "C \ F = {} \ D \ F = {}" using connected_closed [of "F"] \e>0\ not unfolding C_def D_def by (metis Elementary_Metric_Spaces.open_ball F_def closed_cball connected_connected_component inf_bot_left open_closed) moreover have "C \ F \ {}" unfolding disjoint_iff by (metis FCD ComplD image_eqI mem_Collect_eq subsetD x D_def F_def Un_iff \0 < e\ centre_in_ball connected_component_refl_eq) moreover have "D \ F \ {}" unfolding disjoint_iff by (metis ComplI D_def ccs dist_norm ele mem_ball norm_minus_commute not not_le that(1)) ultimately show ?thesis by metis qed moreover have "connected_component_set (f ` S) (f x) \ f ` S" by (auto simp: connected_component_in) ultimately have "connected_component_set (f ` S) (f x) = {f x}" by (auto simp: x F_def) } with assms show ?thesis by blast qed lemma continuous_disconnected_range_constant_eq: "(connected S \ (\f::'a::topological_space \ 'b::real_normed_algebra_1. \t. continuous_on S f \ f ` S \ t \ (\y \ t. connected_component_set t y = {y}) \ f constant_on S))" (is ?thesis1) and continuous_discrete_range_constant_eq: "(connected S \ (\f::'a::topological_space \ 'b::real_normed_algebra_1. continuous_on S f \ (\x \ S. \e. 0 < e \ (\y. y \ S \ (f y \ f x) \ e \ norm(f y - f x))) \ f constant_on S))" (is ?thesis2) and continuous_finite_range_constant_eq: "(connected S \ (\f::'a::topological_space \ 'b::real_normed_algebra_1. continuous_on S f \ finite (f ` S) \ f constant_on S))" (is ?thesis3) proof - have *: "\s t u v. \s \ t; t \ u; u \ v; v \ s\ \ (s \ t) \ (s \ u) \ (s \ v)" by blast have "?thesis1 \ ?thesis2 \ ?thesis3" apply (rule *) using continuous_disconnected_range_constant apply metis apply clarify apply (frule discrete_subset_disconnected; blast) apply (blast dest: finite_implies_discrete) apply (blast intro!: finite_range_constant_imp_connected) done then show ?thesis1 ?thesis2 ?thesis3 by blast+ qed lemma continuous_discrete_range_constant: fixes f :: "'a::topological_space \ 'b::real_normed_algebra_1" assumes S: "connected S" and "continuous_on S f" and "\x. x \ S \ \e>0. \y. y \ S \ f y \ f x \ e \ norm (f y - f x)" shows "f constant_on S" using continuous_discrete_range_constant_eq [THEN iffD1, OF S] assms by blast lemma continuous_finite_range_constant: fixes f :: "'a::topological_space \ 'b::real_normed_algebra_1" assumes "connected S" and "continuous_on S f" and "finite (f ` S)" shows "f constant_on S" using assms continuous_finite_range_constant_eq by blast end diff --git a/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy b/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy --- a/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy +++ b/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy @@ -1,2814 +1,2815 @@ section \Cauchy's Integral Formula\ theory Cauchy_Integral_Formula imports Winding_Numbers begin subsection\Proof\ lemma Cauchy_integral_formula_weak: assumes S: "convex S" and "finite k" and conf: "continuous_on S f" and fcd: "(\x. x \ interior S - k \ f field_differentiable at x)" and z: "z \ interior S - k" and vpg: "valid_path \" and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - let ?fz = "\w. (f w - f z)/(w - z)" obtain f' where f': "(f has_field_derivative f') (at z)" using fcd [OF z] by (auto simp: field_differentiable_def) have pas: "path_image \ \ S" and znotin: "z \ path_image \" using pasz by blast+ have c: "continuous (at x within S) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ S" for x proof (cases "x = z") case True then show ?thesis using LIM_equal [of "z" ?fz "\w. if w = z then f' else ?fz w"] has_field_derivativeD [OF f'] by (force simp add: continuous_within Lim_at_imp_Lim_at_within) next case False then have dxz: "dist x z > 0" by auto have cf: "continuous (at x within S) f" using conf continuous_on_eq_continuous_within that by blast have "continuous (at x within S) (\w. (f w - f z) / (w - z))" by (rule cf continuous_intros | simp add: False)+ then show ?thesis apply (rule continuous_transform_within [OF _ dxz that, of ?fz]) apply (force simp: dist_commute) done qed have fink': "finite (insert z k)" using \finite k\ by blast have *: "((\w. if w = z then f' else ?fz w) has_contour_integral 0) \" proof (rule Cauchy_theorem_convex [OF _ S fink' _ vpg pas loop]) show "(\w. if w = z then f' else ?fz w) field_differentiable at w" if "w \ interior S - insert z k" for w proof (rule field_differentiable_transform_within) show "(\w. ?fz w) field_differentiable at w" using that by (intro derivative_intros fcd; simp) qed (use that in \auto simp add: dist_pos_lt dist_commute\) qed (use c in \force simp: continuous_on_eq_continuous_within\) show ?thesis apply (rule has_contour_integral_eq) using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] apply (auto simp: ac_simps divide_simps) done qed theorem Cauchy_integral_formula_convex_simple: assumes "convex S" and holf: "f holomorphic_on S" and "z \ interior S" "valid_path \" "path_image \ \ S - {z}" "pathfinish \ = pathstart \" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - have "\x. x \ interior S \ f field_differentiable at x" using holf at_within_interior holomorphic_onD interior_subset by fastforce then show ?thesis using assms by (intro Cauchy_integral_formula_weak [where k = "{}"]) (auto simp: holomorphic_on_imp_continuous_on) qed text\ Hence the Cauchy formula for points inside a circle.\ theorem Cauchy_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r" shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) (circlepath z r)" proof - have "r > 0" using assms le_less_trans norm_ge_zero by blast have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) (circlepath z r)" proof (rule Cauchy_integral_formula_weak [where S = "cball z r" and k = "{}"]) show "\x. x \ interior (cball z r) - {} \ f field_differentiable at x" using holf holomorphic_on_imp_differentiable_at by auto have "w \ sphere z r" by simp (metis dist_commute dist_norm not_le order_refl wz) then show "path_image (circlepath z r) \ cball z r - {w}" using \r > 0\ by (auto simp add: cball_def sphere_def) qed (use wz in \simp_all add: dist_norm norm_minus_commute contf\) then show ?thesis by (simp add: winding_number_circlepath assms) qed corollary\<^marker>\tag unimportant\ Cauchy_integral_circlepath_simple: assumes "f holomorphic_on cball z r" "norm(w - z) < r" shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) (circlepath z r)" using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) subsection\<^marker>\tag unimportant\ \General stepping result for derivative formulas\ lemma Cauchy_next_derivative: assumes "continuous_on (path_image \) f'" and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" and int: "\w. w \ S - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" and k: "k \ 0" and "open S" and \: "valid_path \" and w: "w \ S - path_image \" shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) (at w)" (is "?thes2") proof - have "open (S - path_image \)" using \open S\ closed_valid_path_image \ by blast then obtain d where "d>0" and d: "ball w d \ S - path_image \" using w using open_contains_ball by blast have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" by (metis norm_of_nat of_nat_Suc) have cint: "\x. \x \ w; cmod (x - w) < d\ \ (\z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \" using int w d apply (intro contour_integrable_div contour_integrable_diff has_contour_integral_integrable) by (force simp: dist_norm norm_minus_commute) have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) contour_integrable_on \" unfolding eventually_at apply (rule_tac x=d in exI) apply (simp add: \d > 0\ dist_norm field_simps cint) done have bim_g: "bounded (image f' (path_image \))" by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" by (force simp: bounded_pos path_image_def) have twom: "\\<^sub>F n in at w. \x\path_image \. cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" if "0 < e" for e proof - have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" and uw_less: "cmod (u - w) < e * (d/2) ^ (k+2) / (1 + real k)" for u x proof - define ff where [abs_def]: "ff n w = (if n = 0 then inverse(x - w)^k else if n = 1 then k / (x - w)^(Suc k) else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d/2))" if "z \ ball w (d/2)" "i \ 1" for i z proof - have "z \ path_image \" using \x \ path_image \\ d that ball_divide_subset_numeral by blast then have xz[simp]: "x \ z" using \x \ path_image \\ by blast then have neq: "x * x + z * z \ x * (z * 2)" by (blast intro: dest!: sum_sqs_eq) with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" by (simp add: algebra_simps) show ?thesis using \i \ 1\ apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ done qed { fix a::real and b::real assume ab: "a > 0" "b > 0" then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" by (subst mult_le_cancel_left_pos) (use \k \ 0\ in \auto simp: divide_simps\) with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" by (simp add: field_simps) } note canc = this have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d/2) ^ (k + 2)" if "v \ ball w (d/2)" for v proof - have lessd: "\z. cmod (\ z - v) < d/2 \ cmod (w - \ z) < d" by (metis that norm_minus_commute norm_triangle_half_r dist_norm mem_ball) have "d/2 \ cmod (x - v)" using d x that using lessd d x by (auto simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps) then have "d \ cmod (x - v) * 2" by (simp add: field_split_simps) then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" using \0 < d\ order_less_imp_le power_mono by blast have "x \ v" using that using \x \ path_image \\ ball_divide_subset_numeral d by fastforce then show ?thesis using \d > 0\ apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) using dpow_le apply (simp add: field_split_simps) done qed have ub: "u \ ball w (d/2)" using uwd by (simp add: dist_commute dist_norm) have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d/2) ^ k))" using complex_Taylor [OF _ ff1 ff2 _ ub, of w, simplified] by (simp add: ff_def \0 < d\) then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" by (simp add: field_simps) then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) / (cmod (u - w) * real k) \ (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) also have "\ < e" using uw_less \0 < d\ by (simp add: mult_ac divide_simps) finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) / cmod ((u - w) * real k) < e" by (simp add: norm_mult) have "x \ u" using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) show ?thesis apply (rule le_less_trans [OF _ e]) using \k \ 0\ \x \ u\ \u \ w\ apply (simp add: field_simps norm_divide [symmetric]) done qed show ?thesis unfolding eventually_at apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) done qed have 2: "uniform_limit (path_image \) (\n x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) (\x. f' x / (x - w) ^ Suc k) (at w)" unfolding uniform_limit_iff dist_norm proof clarify fix e::real assume "0 < e" have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" and x: "0 \ x" "x \ 1" for u x proof (cases "(f' (\ x)) = 0") case True then show ?thesis by (simp add: \0 < e\) next case False have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k))" by (simp add: field_simps) also have "\ = cmod (f' (\ x)) * cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k)" by (simp add: norm_mult) also have "\ < cmod (f' (\ x)) * (e/C)" using False mult_strict_left_mono [OF ec] by force also have "\ \ e" using C by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) finally show ?thesis . qed show "\\<^sub>F n in at w. \x\path_image \. cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" using twom [OF divide_pos_pos [OF \0 < e\ \C > 0\]] unfolding path_image_def by (force intro: * elim: eventually_mono) qed show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = (f u - f w) / (u - w) / k" if "dist u w < d" for u proof - have u: "u \ S - path_image \" by (metis subsetD d dist_commute mem_ball that) have \
: "((\x. f' x * inverse (x - u) ^ k) has_contour_integral f u) \" "((\x. f' x * inverse (x - w) ^ k) has_contour_integral f w) \" using u w by (simp_all add: field_simps int) show ?thesis apply (rule contour_integral_unique) apply (simp add: diff_divide_distrib algebra_simps \
has_contour_integral_diff has_contour_integral_div) done qed show ?thes2 apply (simp add: has_field_derivative_iff del: power_Suc) apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) apply (simp add: \k \ 0\ **) done qed lemma Cauchy_next_derivative_circlepath: assumes contf: "continuous_on (path_image (circlepath z r)) f" and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" and k: "k \ 0" and w: "w \ ball z r" shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" (is "?thes1") and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" (is "?thes2") proof - have "r > 0" using w using ball_eq_empty by fastforce have wim: "w \ ball z r - path_image (circlepath z r)" using w by (auto simp: dist_norm) show ?thes1 ?thes2 by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; auto simp: vector_derivative_circlepath norm_mult)+ qed text\ In particular, the first derivative formula.\ lemma Cauchy_derivative_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" (is "?thes1") and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" (is "?thes2") proof - have [simp]: "r \ 0" using w using ball_eq_empty by fastforce have f: "continuous_on (path_image (circlepath z r)) f" by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def) have int: "\w. dist z w < r \ ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) show ?thes1 apply (simp add: power2_eq_square) apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) apply (blast intro: int) done have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" apply (simp add: power2_eq_square) apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) apply (blast intro: int) done then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) show ?thes2 by simp (rule fder) qed subsection\Existence of all higher derivatives\ proposition derivative_is_holomorphic: assumes "open S" and fder: "\z. z \ S \ (f has_field_derivative f' z) (at z)" shows "f' holomorphic_on S" proof - have *: "\h. (f' has_field_derivative h) (at z)" if "z \ S" for z proof - obtain r where "r > 0" and r: "cball z r \ S" using open_contains_cball \z \ S\ \open S\ by blast then have holf_cball: "f holomorphic_on cball z r" unfolding holomorphic_on_def using field_differentiable_at_within field_differentiable_def fder by fastforce then have "continuous_on (path_image (circlepath z r)) f" using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" by (auto intro: continuous_intros)+ have contf_cball: "continuous_on (cball z r) f" using holf_cball by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) have holf_ball: "f holomorphic_on ball z r" using holf_cball using ball_subset_cball holomorphic_on_subset by blast { fix w assume w: "w \ ball z r" have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) (at w)" by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) (circlepath z r)" by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) (circlepath z r)" by (simp add: algebra_simps) then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" by (simp add: f'_eq) } note * = this show ?thesis using Cauchy_next_derivative_circlepath [OF contfpi, of 2 f'] \0 < r\ * using centre_in_ball mem_ball by force qed show ?thesis by (simp add: holomorphic_on_open [OF \open S\] *) qed lemma holomorphic_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv f) holomorphic_on S" by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) lemma analytic_deriv [analytic_intros]: "f analytic_on S \ (deriv f) analytic_on S" using analytic_on_holomorphic holomorphic_deriv by auto lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv ^^ n) f holomorphic_on S" by (induction n) (auto simp: holomorphic_deriv) lemma analytic_higher_deriv [analytic_intros]: "f analytic_on S \ (deriv ^^ n) f analytic_on S" unfolding analytic_on_def using holomorphic_higher_deriv by blast lemma has_field_derivative_higher_deriv: "\f holomorphic_on S; open S; x \ S\ \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) lemma valid_path_compose_holomorphic: assumes "valid_path g" and holo:"f holomorphic_on S" and "open S" "path_image g \ S" shows "valid_path (f \ g)" proof (rule valid_path_compose[OF \valid_path g\]) fix x assume "x \ path_image g" then show "f field_differentiable at x" using analytic_on_imp_differentiable_at analytic_on_open assms holo by blast next have "deriv f holomorphic_on S" using holomorphic_deriv holo \open S\ by auto then show "continuous_on (path_image g) (deriv f)" using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto qed subsection\Morera's theorem\ lemma Morera_local_triangle_ball: assumes "\z. z \ S \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ (\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" shows "f analytic_on S" proof - { fix z assume "z \ S" with assms obtain e a where "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" and 0: "\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" by blast have az: "dist a z < e" using mem_ball z by blast have "\e>0. f holomorphic_on ball z e" proof (intro exI conjI) show "f holomorphic_on ball z (e - dist a z)" proof (rule holomorphic_on_subset) show "ball z (e - dist a z) \ ball a e" by (simp add: dist_commute ball_subset_ball_iff) have sub_ball: "\y. dist a y < e \ closed_segment a y \ ball a e" by (meson \0 < e\ centre_in_ball convex_ball convex_contains_segment mem_ball) show "f holomorphic_on ball a e" using triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a] derivative_is_holomorphic[OF open_ball] by (force simp add: 0 \0 < e\ sub_ball) qed qed (simp add: az) } then show ?thesis by (simp add: analytic_on_def) qed lemma Morera_local_triangle: assumes "\z. z \ S \ \t. open t \ z \ t \ continuous_on t f \ (\a b c. convex hull {a,b,c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" shows "f analytic_on S" proof - { fix z assume "z \ S" with assms obtain t where "open t" and z: "z \ t" and contf: "continuous_on t f" and 0: "\a b c. convex hull {a,b,c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" by force then obtain e where "e>0" and e: "ball z e \ t" using open_contains_ball by blast have [simp]: "continuous_on (ball z e) f" using contf using continuous_on_subset e by blast have eq0: "\b c. closed_segment b c \ ball z e \ contour_integral (linepath z b) f + contour_integral (linepath b c) f + contour_integral (linepath c z) f = 0" by (meson 0 z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) have "\e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ (\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" using \e > 0\ eq0 by force } then show ?thesis by (simp add: Morera_local_triangle_ball) qed proposition Morera_triangle: "\continuous_on S f; open S; \a b c. convex hull {a,b,c} \ S \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0\ \ f analytic_on S" using Morera_local_triangle by blast subsection\Combining theorems for higher derivatives including Leibniz rule\ lemma higher_deriv_linear [simp]: "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" by (induction n) auto lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" by (induction n) auto lemma higher_deriv_ident [simp]: "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" proof (induction n) case (Suc n) then show ?case by (metis higher_deriv_linear lambda_one) qed auto lemma higher_deriv_id [simp]: "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" by (simp add: id_def) lemma has_complex_derivative_funpow_1: "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" proof (induction n) case 0 then show ?case by (simp add: id_def) next case (Suc n) then show ?case by (metis DERIV_chain funpow_Suc_right mult.right_neutral) qed lemma higher_deriv_uminus: assumes "f holomorphic_on S" "open S" and z: "z \ S" shows "(deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" using Suc.prems assms has_field_derivative_higher_deriv by auto have "\x. x \ S \ - (deriv ^^ n) f x = (deriv ^^ n) (\w. - f w) x" by (auto simp add: Suc) then have "((deriv ^^ n) (\w. - f w) has_field_derivative - deriv ((deriv ^^ n) f) z) (at z)" using has_field_derivative_transform_within_open [of "\w. -((deriv ^^ n) f w)"] using "*" DERIV_minus Suc.prems \open S\ by blast then show ?case by (simp add: DERIV_imp_deriv) qed lemma higher_deriv_add: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" using Suc.prems assms has_field_derivative_higher_deriv by auto have "\x. x \ S \ (deriv ^^ n) f x + (deriv ^^ n) g x = (deriv ^^ n) (\w. f w + g w) x" by (auto simp add: Suc) then have "((deriv ^^ n) (\w. f w + g w) has_field_derivative deriv ((deriv ^^ n) f) z + deriv ((deriv ^^ n) g) z) (at z)" using has_field_derivative_transform_within_open [of "\w. (deriv ^^ n) f w + (deriv ^^ n) g w"] using "*" Deriv.field_differentiable_add Suc.prems \open S\ by blast then show ?case by (simp add: DERIV_imp_deriv) qed lemma higher_deriv_diff: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" "z \ S" shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" unfolding diff_conv_add_uminus higher_deriv_add using assms higher_deriv_add higher_deriv_uminus holomorphic_on_minus by presburger lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" by (cases k) simp_all lemma higher_deriv_mult: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" shows "(deriv ^^ n) (\w. f w * g w) z = (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have *: "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" "\n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" using Suc.prems assms has_field_derivative_higher_deriv by auto have sumeq: "(\i = 0..n. 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)) = g z * deriv ((deriv ^^ n) f) z + (\i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" apply (simp add: bb algebra_simps sum.distrib) apply (subst (4) sum_Suc_reindex) apply (auto simp: algebra_simps Suc_diff_le intro: sum.cong) done have "((deriv ^^ n) (\w. f w * g w) has_field_derivative (\i = 0..Suc n. (Suc n choose i) * (deriv ^^ i) f z * (deriv ^^ (Suc n - i)) g z)) (at z)" apply (rule has_field_derivative_transform_within_open [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)" _ _ S]) apply (simp add: algebra_simps) apply (rule derivative_eq_intros | simp)+ apply (auto intro: DERIV_mult * \open S\ Suc.prems Suc.IH [symmetric]) by (metis (no_types, lifting) mult.commute sum.cong sumeq) then show ?case unfolding funpow.simps o_apply by (simp add: DERIV_imp_deriv) qed lemma higher_deriv_transform_within_open: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" and fg: "\w. w \ S \ f w = g w" shows "(deriv ^^ i) f z = (deriv ^^ i) g z" using z by (induction i arbitrary: z) (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) lemma higher_deriv_compose_linear: fixes z::complex assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \ S" and fg: "\w. w \ S \ u * w \ T" shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have holo0: "f holomorphic_on (*) u ` S" by (meson fg f holomorphic_on_subset image_subset_iff) have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S" by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T) have holo3: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S" by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros) have "(*) u holomorphic_on S" "f holomorphic_on (*) u ` S" by (rule holo0 holomorphic_intros)+ then have holo1: "(\w. f (u * w)) holomorphic_on S" by (rule holomorphic_on_compose [where g=f, unfolded o_def]) have "deriv ((deriv ^^ n) (\w. f (u * w))) z = deriv (\z. u^n * (deriv ^^ n) f (u*z)) z" proof (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems]) show "(deriv ^^ n) (\w. f (u * w)) holomorphic_on S" by (rule holomorphic_higher_deriv [OF holo1 S]) qed (simp add: Suc.IH) also have "\ = u^n * deriv (\z. (deriv ^^ n) f (u * z)) z" proof - have "(deriv ^^ n) f analytic_on T" by (simp add: analytic_on_open f holomorphic_higher_deriv T) then have "(\w. (deriv ^^ n) f (u * w)) analytic_on S" proof - have "(deriv ^^ n) f \ (*) u holomorphic_on S" by (simp add: holo2 holomorphic_on_compose) then show ?thesis by (simp add: S analytic_on_open o_def) qed then show ?thesis by (intro deriv_cmult analytic_on_imp_differentiable_at [OF _ Suc.prems]) qed also have "\ = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" proof - have "(deriv ^^ n) f field_differentiable at (u * z)" using Suc.prems T f fg holomorphic_higher_deriv holomorphic_on_imp_differentiable_at by blast then show ?thesis by (simp add: deriv_compose_linear) qed finally show ?case by simp qed lemma higher_deriv_add_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_add show ?thesis by (auto simp: analytic_at_two) qed lemma higher_deriv_diff_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_diff show ?thesis by (auto simp: analytic_at_two) qed lemma higher_deriv_uminus_at: "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" using higher_deriv_uminus by (auto simp: analytic_at) lemma higher_deriv_mult_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w * g w) z = (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_mult show ?thesis by (auto simp: analytic_at_two) qed text\ Nonexistence of isolated singularities and a stronger integral formula.\ proposition no_isolated_singularity: fixes z::complex assumes f: "continuous_on S f" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" shows "f holomorphic_on S" proof - { fix z assume "z \ S" and cdf: "\x. x \ S - K \ f field_differentiable at x" have "f field_differentiable at z" proof (cases "z \ K") case False then show ?thesis by (blast intro: cdf \z \ S\) next case True with finite_set_avoid [OF K, of z] obtain d where "d>0" and d: "\x. \x\K; x \ z\ \ d \ dist z x" by blast obtain e where "e>0" and e: "ball z e \ S" using S \z \ S\ by (force simp: open_contains_ball) have fde: "continuous_on (ball z (min d e)) f" by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) have cont: "{a,b,c} \ ball z (min d e) \ continuous_on (convex hull {a, b, c}) f" for a b c by (simp add: hull_minimal continuous_on_subset [OF fde]) have fd: "\{a,b,c} \ ball z (min d e); x \ interior (convex hull {a, b, c}) - K\ \ f field_differentiable at x" for a b c x by (metis cdf Diff_iff Int_iff ball_min_Int subsetD convex_ball e interior_mono interior_subset subset_hull) obtain g where "\w. w \ ball z (min d e) \ (g has_field_derivative f w) (at w within ball z (min d e))" apply (rule contour_integral_convex_primitive [OF convex_ball fde Cauchy_theorem_triangle_cofinite [OF _ K]]) using cont fd by auto then have "f holomorphic_on ball z (min d e)" by (metis open_ball at_within_open derivative_is_holomorphic) then show ?thesis unfolding holomorphic_on_def by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) qed } with holf S K show ?thesis by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) qed lemma no_isolated_singularity': fixes z::complex assumes f: "\z. z \ K \ (f \ f z) (at z within S)" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" shows "f holomorphic_on S" proof (rule no_isolated_singularity[OF _ assms(2-)]) show "continuous_on S f" unfolding continuous_on_def proof fix z assume z: "z \ S" show "(f \ f z) (at z within S)" proof (cases "z \ K") case False from holf have "continuous_on (S - K) f" by (rule holomorphic_on_imp_continuous_on) with z False have "(f \ f z) (at z within (S - K))" by (simp add: continuous_on_def) also from z K S False have "at z within (S - K) = at z within S" by (subst (1 2) at_within_open) (auto intro: finite_imp_closed) finally show "(f \ f z) (at z within S)" . qed (insert assms z, simp_all) qed qed proposition Cauchy_integral_formula_convex: assumes S: "convex S" and K: "finite K" and contf: "continuous_on S f" and fcd: "(\x. x \ interior S - K \ f field_differentiable at x)" and z: "z \ interior S" and vpg: "valid_path \" and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - have *: "\x. x \ interior S \ f field_differentiable at x" unfolding holomorphic_on_open [symmetric] field_differentiable_def using no_isolated_singularity [where S = "interior S"] by (meson K contf continuous_at_imp_continuous_on continuous_on_interior fcd field_differentiable_at_within field_differentiable_def holomorphic_onI holomorphic_on_imp_differentiable_at open_interior) show ?thesis by (rule Cauchy_integral_formula_weak [OF S finite.emptyI contf]) (use * assms in auto) qed text\ Formula for higher derivatives.\ lemma Cauchy_has_contour_integral_higher_derivative_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) (circlepath z r)" using w proof (induction k arbitrary: w) case 0 then show ?case using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) next case (Suc k) have [simp]: "r > 0" using w using ball_eq_empty by fastforce have f: "continuous_on (path_image (circlepath z r)) f" by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def less_imp_le) obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] by (auto simp: contour_integrable_on_def) then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" by (rule contour_integral_unique) have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" using Suc.prems assms has_field_derivative_higher_deriv by auto then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" by (force simp: field_differentiable_def) have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) also have "\ = of_nat (Suc k) * X" by (simp only: con) finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" by (metis deriv_cmult dnf_diff) then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" by (simp add: field_simps) then show ?case using of_nat_eq_0_iff X by fastforce qed lemma Cauchy_higher_derivative_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" (is "?thes1") and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" (is "?thes2") proof - have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) (circlepath z r)" using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] by simp show ?thes1 using * using contour_integrable_on_def by blast show ?thes2 unfolding contour_integral_unique [OF *] by (simp add: field_split_simps) qed corollary Cauchy_contour_integral_circlepath: assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) lemma Cauchy_contour_integral_circlepath_2: assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" using Cauchy_contour_integral_circlepath [OF assms, of 1] by (simp add: power2_eq_square) subsection\A holomorphic function is analytic, i.e. has local power series\ theorem holomorphic_power_series: assumes holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" proof - \ \Replacing \<^term>\r\ and the original (weak) premises with stronger ones\ obtain r where "r > 0" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" proof have "cball z ((r + dist w z) / 2) \ ball z r" using w by (simp add: dist_commute field_sum_of_halves subset_eq) then show "f holomorphic_on cball z ((r + dist w z) / 2)" by (rule holomorphic_on_subset [OF holf]) have "r > 0" using w by clarsimp (metis dist_norm le_less_trans norm_ge_zero) then show "0 < (r + dist w z) / 2" by simp (use zero_le_dist [of w z] in linarith) qed (use w in \auto simp: dist_commute\) then have holf: "f holomorphic_on ball z r" using ball_subset_cball holomorphic_on_subset by blast have contf: "continuous_on (cball z r) f" by (simp add: holfc holomorphic_on_imp_continuous_on) have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" by (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) (simp add: \0 < r\) obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" proof show "\u. cmod (u - z) = r \ r - dist z w \ cmod (u - w)" by (metis add_diff_eq diff_add_cancel dist_norm norm_diff_ineq) qed (use w in \auto simp: dist_norm norm_minus_commute\) have ul: "uniform_limit (sphere z r) (\n x. (\kx. f x / (x - w)) sequentially" unfolding uniform_limit_iff dist_norm proof clarify fix e::real assume "0 < e" have rr: "0 \ (r - k) / r" "(r - k) / r < 1" using k by auto obtain n where n: "((r - k) / r) ^ n < e / B * k" using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force have "norm ((\k N" and r: "r = dist z u" for N u proof - have N: "((r - k) / r) ^ N < e / B * k" using le_less_trans [OF power_decreasing n] using \n \ N\ k by auto have u [simp]: "(u \ z) \ (u \ w)" using \0 < r\ r w by auto have wzu_not1: "(w - z) / (u - z) \ 1" by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) have "norm ((\kk = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)" using \0 < B\ apply (auto simp: geometric_sum [OF wzu_not1]) apply (simp add: field_simps norm_mult [symmetric]) done also have "\ = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) also have "\ = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" by (simp add: algebra_simps) also have "\ = norm (w - z) ^ N * norm (f u) / r ^ N" by (simp add: norm_mult norm_power norm_minus_commute) also have "\ \ (((r - k)/r)^N) * B" using \0 < r\ w k by (simp add: B divide_simps mult_mono r wz_eq) also have "\ < e * k" using \0 < B\ N by (simp add: divide_simps) also have "\ \ e * norm (u - w)" using r kle \0 < e\ by (simp add: dist_commute dist_norm) finally show ?thesis by (simp add: field_split_simps norm_divide del: power_Suc) qed with \0 < r\ show "\\<^sub>F n in sequentially. \x\sphere z r. norm ((\k: "\x k. k\ {.. (\u. (w - z) ^ k * (f u / (u - z) ^ Suc k)) contour_integrable_on circlepath z r" using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] by (simp add: field_simps) have eq: "\\<^sub>F x in sequentially. contour_integral (circlepath z r) (\u. \kku. f u / (u - z) ^ Suc k) * (w - z) ^ k)" apply (rule eventuallyI) apply (subst contour_integral_sum, simp) apply (simp_all only: \
contour_integral_lmul cint algebra_simps) done have "\u k. k \ {.. (\x. f x / (x - z) ^ Suc k) contour_integrable_on circlepath z r" using \0 < r\ by (force intro!: Cauchy_higher_derivative_integral_circlepath [OF contf holf]) then have "\u. (\y. \kk. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) sums contour_integral (circlepath z r) (\u. f u/(u - w))" unfolding sums_def using \0 < r\ by (intro Lim_transform_eventually [OF _ eq] contour_integral_uniform_limit_circlepath [OF eventuallyI ul]) auto then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) sums (2 * of_real pi * \ * f w)" using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" by (rule sums_divide) then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) sums f w" by (simp add: field_simps) then show ?thesis by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) qed subsection\The Liouville theorem and the Fundamental Theorem of Algebra\ text\ These weak Liouville versions don't even need the derivative formula.\ lemma Liouville_weak_0: assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" shows "f z = 0" proof (rule ccontr) assume fz: "f z \ 0" with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" by (auto simp: dist_norm) define R where "R = 1 + \B\ + norm z" have "R > 0" unfolding R_def proof - have "0 \ cmod z + \B\" by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) then show "0 < 1 + \B\ + cmod z" by linarith qed have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" using continuous_on_subset holf holomorphic_on_subset \0 < R\ by (force intro: holomorphic_on_imp_continuous_on Cauchy_integral_circlepath) have "cmod (x - z) = R \ cmod (f x) * 2 < cmod (f z)" for x unfolding R_def by (rule B) (use norm_triangle_ineq4 [of x z] in auto) with \R > 0\ fz show False using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] by (auto simp: less_imp_le norm_mult norm_divide field_split_simps) qed proposition Liouville_weak: assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" shows "f z = l" using Liouville_weak_0 [of "\z. f z - l"] by (simp add: assms holomorphic_on_diff LIM_zero) proposition Liouville_weak_inverse: assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" obtains z where "f z = 0" proof - { assume f: "\z. f z \ 0" have 1: "(\x. 1 / f x) holomorphic_on UNIV" by (simp add: holomorphic_on_divide assms f) have 2: "((\x. 1 / f x) \ 0) at_infinity" proof (rule tendstoI [OF eventually_mono]) fix e::real assume "e > 0" show "eventually (\x. 2/e \ cmod (f x)) at_infinity" by (rule_tac B="2/e" in unbounded) qed (simp add: dist_norm norm_divide field_split_simps) have False using Liouville_weak_0 [OF 1 2] f by simp } then show ?thesis using that by blast qed text\ In particular we get the Fundamental Theorem of Algebra.\ theorem fundamental_theorem_of_algebra: fixes a :: "nat \ complex" assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" obtains z where "(\i\n. a i * z^i) = 0" using assms proof (elim disjE bexE) assume "a 0 = 0" then show ?thesis by (auto simp: that [of 0]) next fix i assume i: "i \ {1..n}" and nz: "a i \ 0" have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" by (rule holomorphic_intros)+ show thesis proof (rule Liouville_weak_inverse [OF 1]) show "\\<^sub>F x in at_infinity. B \ cmod (\i\n. a i * x ^ i)" for B using i nz by (intro polyfun_extremal exI[of _ i]) auto qed (use that in auto) qed subsection\Weierstrass convergence theorem\ lemma holomorphic_uniform_limit: assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" and ulim: "uniform_limit (cball z r) f g F" and F: "\ trivial_limit F" obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (force simp: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) next case equal then show ?thesis by (force simp: holomorphic_on_def intro: that) next case greater have contg: "continuous_on (cball z r) g" using cont uniform_limit_theorem [OF eventually_mono ulim F] by blast have "path_image (circlepath z r) \ cball z r" using \0 < r\ by auto then have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" by (intro continuous_intros continuous_on_subset [OF contg]) have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" if w: "w \ ball z r" for w proof - define d where "d = (r - norm(w - z))" have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" using w by (auto intro: eventually_mono [OF cont] Cauchy_higher_derivative_integral_circlepath [where k=0, simplified]) have "\e. \0 < r; 0 < d; 0 < e\ \ \\<^sub>F n in F. \x\sphere z r. x \ w \ cmod (f n x - g x) < e * cmod (x - w)" apply (rule_tac e1="e * d" in eventually_mono [OF uniform_limitD [OF ulim]]) apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+ done then have ul_less: "uniform_limit (sphere z r) (\n x. f n x / (x - w)) (\x. g x / (x - w)) F" using greater \0 < d\ by (auto simp add: uniform_limit_iff dist_norm norm_divide diff_divide_distrib [symmetric] divide_simps) have g_cint: "(\u. g u/(u - w)) contour_integrable_on circlepath z r" by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) have cif_tends_cig: "((\n. contour_integral(circlepath z r) (\u. f n u / (u - w))) \ contour_integral(circlepath z r) (\u. g u/(u - w))) F" by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) have f_tends_cig: "((\n. 2 * of_real pi * \ * f n w) \ contour_integral (circlepath z r) (\u. g u / (u - w))) F" proof (rule Lim_transform_eventually) show "\\<^sub>F x in F. contour_integral (circlepath z r) (\u. f x u / (u - w)) = 2 * of_real pi * \ * f x w" using w\0 < d\ d_def by (auto intro: eventually_mono [OF cont contour_integral_unique [OF Cauchy_integral_circlepath]]) qed (auto simp: cif_tends_cig) have "\e. 0 < e \ \\<^sub>F n in F. dist (f n w) (g w) < e" by (rule eventually_mono [OF uniform_limitD [OF ulim]]) (use w in auto) then have "((\n. 2 * of_real pi * \ * f n w) \ 2 * of_real pi * \ * g w) F" by (rule tendsto_mult_left [OF tendstoI]) then have "((\u. g u / (u - w)) has_contour_integral 2 * of_real pi * \ * g w) (circlepath z r)" using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w by fastforce then have "((\u. g u / (2 * of_real pi * \ * (u - w))) has_contour_integral g w) (circlepath z r)" using has_contour_integral_div [where c = "2 * of_real pi * \"] by (force simp: field_simps) then show ?thesis by (simp add: dist_norm) qed show ?thesis using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] by (fastforce simp add: holomorphic_on_open contg intro: that) qed text\ Version showing that the limit is the limit of the derivatives.\ proposition has_complex_derivative_uniform_limit: fixes z::complex assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" and ulim: "uniform_limit (cball z r) f g F" and F: "\ trivial_limit F" and "0 < r" obtains g' where "continuous_on (cball z r) g" "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" proof - let ?conint = "contour_integral (circlepath z r)" have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] ulim F]; auto simp: holomorphic_on_open field_differentiable_def)+ then obtain g' where g': "\x. x \ ball z r \ (g has_field_derivative g' x) (at x)" using DERIV_deriv_iff_has_field_derivative by (fastforce simp add: holomorphic_on_open) then have derg: "\x. x \ ball z r \ deriv g x = g' x" by (simp add: DERIV_imp_deriv) have tends_f'n_g': "((\n. f' n w) \ g' w) F" if w: "w \ ball z r" for w proof - have eq_f': "?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \)" if cont_fn: "continuous_on (cball z r) (f n)" and fnd: "\w. w \ ball z r \ (f n has_field_derivative f' n w) (at w)" for n proof - have hol_fn: "f n holomorphic_on ball z r" using fnd by (force simp: holomorphic_on_open) have "(f n has_field_derivative 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)) (at w)" by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) then have f': "f' n w = 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)" using DERIV_unique [OF fnd] w by blast show ?thesis by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] field_split_simps) qed define d where "d = (r - norm(w - z))^2" have "d > 0" using w by (simp add: dist_commute dist_norm d_def) have dle: "d \ cmod ((y - w)\<^sup>2)" if "r = cmod (z - y)" for y proof - have "w \ ball z (cmod (z - y))" using that w by fastforce then have "cmod (w - z) \ cmod (z - y)" by (simp add: dist_complex_def norm_minus_commute) moreover have "cmod (z - y) - cmod (w - z) \ cmod (y - w)" by (metis diff_add_cancel diff_add_eq_diff_diff_swap norm_minus_commute norm_triangle_ineq2) ultimately show ?thesis using that by (simp add: d_def norm_power power_mono) qed have 1: "\\<^sub>F n in F. (\x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" by (force simp: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) have 2: "uniform_limit (sphere z r) (\n x. f n x / (x - w)\<^sup>2) (\x. g x / (x - w)\<^sup>2) F" unfolding uniform_limit_iff proof clarify fix e::real assume "e > 0" with \r > 0\ have "\\<^sub>F n in F. \x. x \ w \ cmod (z - x) = r \ cmod (f n x - g x) < e * cmod ((x - w)\<^sup>2)" by (force simp: \0 < d\ dist_norm dle intro: less_le_trans eventually_mono [OF uniform_limitD [OF ulim], of "e*d"]) with \r > 0\ \e > 0\ show "\\<^sub>F n in F. \x\sphere z r. dist (f n x / (x - w)\<^sup>2) (g x / (x - w)\<^sup>2) < e" by (simp add: norm_divide field_split_simps sphere_def dist_norm) qed have "((\n. contour_integral (circlepath z r) (\x. f n x / (x - w)\<^sup>2)) \ contour_integral (circlepath z r) ((\x. g x / (x - w)\<^sup>2))) F" by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \0 < r\]) then have tendsto_0: "((\n. 1 / (2 * of_real pi * \) * (?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2))) \ 0) F" using Lim_null by (force intro!: tendsto_mult_right_zero) have "((\n. f' n w - g' w) \ 0) F" apply (rule Lim_transform_eventually [OF tendsto_0]) apply (force simp: divide_simps intro: eq_f' eventually_mono [OF cont]) done then show ?thesis using Lim_null by blast qed obtain g' where "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" by (blast intro: tends_f'n_g' g') then show ?thesis using g using that by blast qed subsection\<^marker>\tag unimportant\ \Some more simple/convenient versions for applications\ lemma holomorphic_uniform_sequence: assumes S: "open S" and hol_fn: "\n. (f n) holomorphic_on S" and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" shows "g holomorphic_on S" proof - have "\f'. (g has_field_derivative f') (at z)" if "z \ S" for z proof - obtain r where "0 < r" and r: "cball z r \ S" and ul: "uniform_limit (cball z r) f g sequentially" using ulim_g [OF \z \ S\] by blast have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ f n holomorphic_on ball z r" proof (intro eventuallyI conjI) show "continuous_on (cball z r) (f x)" for x using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r by blast show "f x holomorphic_on ball z r" for x by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) qed show ?thesis using \0 < r\ centre_in_ball ul by (auto simp: holomorphic_on_open intro: holomorphic_uniform_limit [OF *]) qed with S show ?thesis by (simp add: holomorphic_on_open) qed lemma has_complex_derivative_uniform_sequence: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ ((f n) has_field_derivative f' n x) (at x)" and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" shows "\g'. \x \ S. (g has_field_derivative g' x) (at x) \ ((\n. f' n x) \ g' x) sequentially" proof - have y: "\y. (g has_field_derivative y) (at z) \ (\n. f' n z) \ y" if "z \ S" for z proof - obtain r where "0 < r" and r: "cball z r \ S" and ul: "uniform_limit (cball z r) f g sequentially" using ulim_g [OF \z \ S\] by blast have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))" proof (intro eventuallyI conjI ballI) show "continuous_on (cball z r) (f x)" for x by (meson S continuous_on_subset hfd holomorphic_on_imp_continuous_on holomorphic_on_open r) show "w \ ball z r \ (f x has_field_derivative f' x w) (at w)" for w x using ball_subset_cball hfd r by blast qed show ?thesis by (rule has_complex_derivative_uniform_limit [OF *, of g]) (use \0 < r\ ul in \force+\) qed show ?thesis by (rule bchoice) (blast intro: y) qed subsection\On analytic functions defined by a series\ lemma series_and_derivative_comparison: fixes S :: "complex set" assumes S: "open S" and h: "summable h" and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" and to_g: "\\<^sub>F n in sequentially. \x\S. norm (f n x) \ h n" obtains g g' where "\x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" proof - obtain g where g: "uniform_limit S (\n x. \id>0. cball x d \ S \ uniform_limit (cball x d) (\n x. \i S" for x proof - obtain d where "d>0" and d: "cball x d \ S" using open_contains_cball [of "S"] \x \ S\ S by blast show ?thesis proof (intro conjI exI) show "uniform_limit (cball x d) (\n x. \id > 0\ d in auto) qed have "\x. x \ S \ (\n. \i g x" by (metis tendsto_uniform_limitI [OF g]) moreover have "\g'. \x\S. (g has_field_derivative g' x) (at x) \ (\n. \i g' x" by (rule has_complex_derivative_uniform_sequence [OF S]) (auto intro: * hfd DERIV_sum)+ ultimately show ?thesis by (metis sums_def that) qed text\A version where we only have local uniform/comparative convergence.\ lemma series_and_derivative_comparison_local: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ (\\<^sub>F n in sequentially. \y\ball x d \ S. norm (f n y) \ h n)" shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" proof - have "\y. (\n. f n z) sums (\n. f n z) \ (\n. f' n z) sums y \ ((\x. \n. f n x) has_field_derivative y) (at z)" if "z \ S" for z proof - obtain d h where "0 < d" "summable h" and le_h: "\\<^sub>F n in sequentially. \y\ball z d \ S. norm (f n y) \ h n" using to_g \z \ S\ by meson then obtain r where "r>0" and r: "ball z r \ ball z d \ S" using \z \ S\ S by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) have 1: "open (ball z d \ S)" by (simp add: open_Int S) have 2: "\n x. x \ ball z d \ S \ (f n has_field_derivative f' n x) (at x)" by (auto simp: hfd) obtain g g' where gg': "\x \ ball z d \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" by (auto intro: le_h series_and_derivative_comparison [OF 1 \summable h\ hfd]) then have "(\n. f' n z) sums g' z" by (meson \0 < r\ centre_in_ball contra_subsetD r) moreover have "(\n. f n z) sums (\n. f n z)" using summable_sums centre_in_ball \0 < d\ \summable h\ le_h by (metis (full_types) Int_iff gg' summable_def that) moreover have "((\x. \n. f n x) has_field_derivative g' z) (at z)" proof (rule has_field_derivative_transform_within) show "\x. dist x z < r \ g x = (\n. f n x)" by (metis subsetD dist_commute gg' mem_ball r sums_unique) qed (use \0 < r\ gg' \z \ S\ \0 < d\ in auto) ultimately show ?thesis by auto qed then show ?thesis by (rule_tac x="\x. suminf (\n. f n x)" in exI) meson qed text\Sometimes convenient to compare with a complex series of positive reals. (?)\ lemma series_and_derivative_comparison_complex: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" apply (rule series_and_derivative_comparison_local [OF S hfd], assumption) apply (rule ex_forward [OF to_g], assumption) apply (erule exE) apply (rule_tac x="Re \ h" in exI) apply (force simp: summable_Re o_def nonneg_Reals_cmod_eq_Re image_subset_iff) done text\Sometimes convenient to compare with a complex series of positive reals. (?)\ lemma series_differentiable_comparison_complex: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ f n field_differentiable (at x)" and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" obtains g where "\x \ S. ((\n. f n x) sums g x) \ g field_differentiable (at x)" proof - have hfd': "\n x. x \ S \ (f n has_field_derivative deriv (f n) x) (at x)" using hfd field_differentiable_derivI by blast have "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. deriv (f n) x) sums g' x) \ (g has_field_derivative g' x) (at x)" by (metis series_and_derivative_comparison_complex [OF S hfd' to_g]) then show ?thesis using field_differentiable_def that by blast qed text\In particular, a power series is analytic inside circle of convergence.\ lemma power_series_and_derivative_0: fixes a :: "nat \ complex" and r::real assumes "summable (\n. a n * r^n)" shows "\g g'. \z. cmod z < r \ ((\n. a n * z^n) sums g z) \ ((\n. of_nat n * a n * z^(n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" proof (cases "0 < r") case True have der: "\n z. ((\x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)" by (rule derivative_eq_intros | simp)+ have y_le: "cmod y \ cmod (of_real r + of_real (cmod z)) / 2" if "cmod (z - y) * 2 < r - cmod z" for z y proof - have "cmod y * 2 \ r + cmod z" using norm_triangle_ineq2 [of y z] that by (simp only: diff_le_eq norm_minus_commute mult_2) then show ?thesis using \r > 0\ by (auto simp: algebra_simps norm_mult norm_divide norm_power simp flip: of_real_add) qed have "summable (\n. a n * complex_of_real r ^ n)" using assms \r > 0\ by simp moreover have "\z. cmod z < r \ cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)" using \r > 0\ by (simp flip: of_real_add) ultimately have sum: "\z. cmod z < r \ summable (\n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)" by (rule power_series_conv_imp_absconv_weak) have "\g g'. \z \ ball 0 r. (\n. (a n) * z ^ n) sums g z \ (\n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \ (g has_field_derivative g' z) (at z)" apply (rule series_and_derivative_comparison_complex [OF open_ball der]) apply (rule_tac x="(r - norm z)/2" in exI) apply (rule_tac x="\n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI) using \r > 0\ apply (auto simp: sum eventually_sequentially norm_mult norm_power dist_norm intro!: mult_left_mono power_mono y_le) done then show ?thesis by (simp add: ball_def) next case False then show ?thesis unfolding not_less using less_le_trans norm_not_less_zero by blast qed proposition\<^marker>\tag unimportant\ power_series_and_derivative: fixes a :: "nat \ complex" and r::real assumes "summable (\n. a n * r^n)" obtains g g' where "\z \ ball w r. ((\n. a n * (z - w) ^ n) sums g z) \ ((\n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" using power_series_and_derivative_0 [OF assms] apply clarify apply (rule_tac g="(\z. g(z - w))" in that) using DERIV_shift [where z="-w"] apply (auto simp: norm_minus_commute Ball_def dist_norm) done proposition\<^marker>\tag unimportant\ power_series_holomorphic: assumes "\w. w \ ball z r \ ((\n. a n*(w - z)^n) sums f w)" shows "f holomorphic_on ball z r" proof - have "\f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w proof - have inb: "z + complex_of_real ((dist z w + r) / 2) \ ball z r" proof - have wz: "cmod (w - z) < r" using w by (auto simp: field_split_simps dist_norm norm_minus_commute) then have "0 \ r" by (meson less_eq_real_def norm_ge_zero order_trans) show ?thesis using w by (simp add: dist_norm \0\r\ flip: of_real_add) qed have sum: "summable (\n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))" using assms [OF inb] by (force simp: summable_def dist_norm) obtain g g' where gg': "\u. u \ ball z ((cmod (z - w) + r) / 2) \ (\n. a n * (u - z) ^ n) sums g u \ (\n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \ (g has_field_derivative g' u) (at u)" by (rule power_series_and_derivative [OF sum, of z]) fastforce have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u proof - have less: "cmod (z - u) * 2 < cmod (z - w) + r" using that dist_triangle2 [of z u w] by (simp add: dist_norm [symmetric] algebra_simps) have "(\n. a n * (u - z) ^ n) sums g u" "(\n. a n * (u - z) ^ n) sums f u" using gg' [of u] less w by (auto simp: assms dist_norm) then show ?thesis by (metis sums_unique2) qed have "(f has_field_derivative g' w) (at w)" by (rule has_field_derivative_transform_within [where d="(r - norm(z - w))/2"]) (use w gg' [of w] in \(force simp: dist_norm)+\) then show ?thesis .. qed then show ?thesis by (simp add: holomorphic_on_open) qed corollary holomorphic_iff_power_series: "f holomorphic_on ball z r \ (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" apply (intro iffI ballI holomorphic_power_series, assumption+) apply (force intro: power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) done lemma power_series_analytic: "(\w. w \ ball z r \ (\n. a n*(w - z)^n) sums f w) \ f analytic_on ball z r" by (force simp: analytic_on_open intro!: power_series_holomorphic) lemma analytic_iff_power_series: "f analytic_on ball z r \ (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" by (simp add: analytic_on_open holomorphic_iff_power_series) subsection\<^marker>\tag unimportant\ \Equality between holomorphic functions, on open ball then connected set\ lemma holomorphic_fun_eq_on_ball: "\f holomorphic_on ball z r; g holomorphic_on ball z r; w \ ball z r; \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ \ f w = g w" by (auto simp: holomorphic_iff_power_series sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) lemma holomorphic_fun_eq_0_on_ball: "\f holomorphic_on ball z r; w \ ball z r; \n. (deriv ^^ n) f z = 0\ \ f w = 0" by (auto simp: holomorphic_iff_power_series sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) lemma holomorphic_fun_eq_0_on_connected: assumes holf: "f holomorphic_on S" and "open S" and cons: "connected S" and der: "\n. (deriv ^^ n) f z = 0" and "z \ S" "w \ S" shows "f w = 0" proof - have *: "ball x e \ (\n. {w \ S. (deriv ^^ n) f w = 0})" if "\u. (deriv ^^ u) f x = 0" "ball x e \ S" for x e proof - have "(deriv ^^ m) ((deriv ^^ n) f) x = 0" for m n by (metis funpow_add o_apply that(1)) then have "\x' n. dist x x' < e \ (deriv ^^ n) f x' = 0" using \open S\ by (meson holf holomorphic_fun_eq_0_on_ball holomorphic_higher_deriv holomorphic_on_subset mem_ball that(2)) with that show ?thesis by auto qed obtain e where "e>0" and e: "ball w e \ S" using openE [OF \open S\ \w \ S\] . then have holfb: "f holomorphic_on ball w e" using holf holomorphic_on_subset by blast have "open (\n. {w \ S. (deriv ^^ n) f w = 0})" using \open S\ apply (simp add: open_contains_ball Ball_def) apply (erule all_forward) using "*" by auto blast+ then have "openin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" by (force intro: open_subset) moreover have "closedin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" using assms by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) moreover have "(\n. {w \ S. (deriv ^^ n) f w = 0}) = S \ f w = 0" using \e>0\ e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) ultimately show ?thesis using cons der \z \ S\ by (auto simp add: connected_clopen) qed lemma holomorphic_fun_eq_on_connected: assumes "f holomorphic_on S" "g holomorphic_on S" and "open S" "connected S" and "\n. (deriv ^^ n) f z = (deriv ^^ n) g z" and "z \ S" "w \ S" shows "f w = g w" proof (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" S z, simplified]) show "(\x. f x - g x) holomorphic_on S" by (intro assms holomorphic_intros) show "\n. (deriv ^^ n) (\x. f x - g x) z = 0" using assms higher_deriv_diff by auto qed (use assms in auto) lemma holomorphic_fun_eq_const_on_connected: assumes holf: "f holomorphic_on S" and "open S" and cons: "connected S" and der: "\n. 0 < n \ (deriv ^^ n) f z = 0" and "z \ S" "w \ S" shows "f w = f z" proof (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" S z, simplified]) show "(\w. f w - f z) holomorphic_on S" by (intro assms holomorphic_intros) show "\n. (deriv ^^ n) (\w. f w - f z) z = 0" by (subst higher_deriv_diff) (use assms in \auto intro: holomorphic_intros\) qed (use assms in auto) subsection\<^marker>\tag unimportant\ \Some basic lemmas about poles/singularities\ lemma pole_lemma: assumes holf: "f holomorphic_on S" and a: "a \ interior S" shows "(\z. if z = a then deriv f a else (f z - f a) / (z - a)) holomorphic_on S" (is "?F holomorphic_on S") proof - have F1: "?F field_differentiable (at u within S)" if "u \ S" "u \ a" for u proof - have fcd: "f field_differentiable at u within S" using holf holomorphic_on_def by (simp add: \u \ S\) have cd: "(\z. (f z - f a) / (z - a)) field_differentiable at u within S" by (rule fcd derivative_intros | simp add: that)+ have "0 < dist a u" using that dist_nz by blast then show ?thesis by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \u \ S\) qed have F2: "?F field_differentiable at a" if "0 < e" "ball a e \ S" for e proof - have holfb: "f holomorphic_on ball a e" by (rule holomorphic_on_subset [OF holf \ball a e \ S\]) have 2: "?F holomorphic_on ball a e - {a}" using mem_ball that by (auto simp add: holomorphic_on_def simp flip: field_differentiable_def intro: F1 field_differentiable_within_subset) have "isCont (\z. if z = a then deriv f a else (f z - f a) / (z - a)) x" if "dist a x < e" for x proof (cases "x=a") case True then have "f field_differentiable at a" using holfb \0 < e\ holomorphic_on_imp_differentiable_at by auto with True show ?thesis by (auto simp: continuous_at has_field_derivative_iff simp flip: DERIV_deriv_iff_field_differentiable elim: rev_iffD1 [OF _ LIM_equal]) next case False with 2 that show ?thesis by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) qed then have 1: "continuous_on (ball a e) ?F" by (clarsimp simp: continuous_on_eq_continuous_at) have "?F holomorphic_on ball a e" by (auto intro: no_isolated_singularity [OF 1 2]) with that show ?thesis by (simp add: holomorphic_on_open field_differentiable_def [symmetric] field_differentiable_at_within) qed show ?thesis proof fix x assume "x \ S" show "?F field_differentiable at x within S" proof (cases "x=a") case True then show ?thesis using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) next case False with F1 \x \ S\ show ?thesis by blast qed qed qed lemma pole_theorem: assumes holg: "g holomorphic_on S" and a: "a \ interior S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) holomorphic_on S" using pole_lemma [OF holg a] by (rule holomorphic_transform) (simp add: eq field_split_simps) lemma pole_lemma_open: assumes "f holomorphic_on S" "open S" shows "(\z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on S" proof (cases "a \ S") case True with assms interior_eq pole_lemma show ?thesis by fastforce next case False with assms show ?thesis apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) apply (rule field_differentiable_transform_within [where f = "\z. (f z - f a)/(z - a)" and d = 1]) apply (rule derivative_intros | force)+ done qed lemma pole_theorem_open: assumes holg: "g holomorphic_on S" and S: "open S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) holomorphic_on S" using pole_lemma_open [OF holg S] by (rule holomorphic_transform) (auto simp: eq divide_simps) lemma pole_theorem_0: assumes holg: "g holomorphic_on S" and a: "a \ interior S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f holomorphic_on S" using pole_theorem [OF holg a eq] by (rule holomorphic_transform) (auto simp: eq field_split_simps) lemma pole_theorem_open_0: assumes holg: "g holomorphic_on S" and S: "open S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f holomorphic_on S" using pole_theorem_open [OF holg S eq] by (rule holomorphic_transform) (auto simp: eq field_split_simps) lemma pole_theorem_analytic: assumes g: "g analytic_on S" and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" (is "?F analytic_on S") unfolding analytic_on_def proof fix x assume "x \ S" with g obtain e where "0 < e" and e: "g holomorphic_on ball x e" by (auto simp add: analytic_on_def) obtain d where "0 < d" and d: "\w. w \ ball x d - {a} \ g w = (w - a) * f w" using \x \ S\ eq by blast have "?F holomorphic_on ball x (min d e)" using d e \x \ S\ by (fastforce simp: holomorphic_on_subset subset_ball intro!: pole_theorem_open) then show "\e>0. ?F holomorphic_on ball x e" using \0 < d\ \0 < e\ not_le by fastforce qed lemma pole_theorem_analytic_0: assumes g: "g analytic_on S" and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" and [simp]: "f a = deriv g a" "g a = 0" shows "f analytic_on S" proof - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" by auto show ?thesis using pole_theorem_analytic [OF g eq] by simp qed lemma pole_theorem_analytic_open_superset: assumes g: "g analytic_on S" and "S \ T" "open T" and eq: "\z. z \ T - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" proof (rule pole_theorem_analytic [OF g]) fix z assume "z \ S" then obtain e where "0 < e" and e: "ball z e \ T" using assms openE by blast then show "\d>0. \w\ball z d - {a}. g w = (w - a) * f w" using eq by auto qed lemma pole_theorem_analytic_open_superset_0: assumes g: "g analytic_on S" "S \ T" "open T" "\z. z \ T - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f analytic_on S" proof - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" by auto have "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" by (rule pole_theorem_analytic_open_superset [OF g]) then show ?thesis by simp qed subsection\General, homology form of Cauchy's theorem\ text\Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\ lemma contour_integral_continuous_on_linepath_2D: assumes "open U" and cont_dw: "\w. w \ U \ F w contour_integrable_on (linepath a b)" and cond_uu: "continuous_on (U \ U) (\(x,y). F x y)" and abu: "closed_segment a b \ U" shows "continuous_on U (\w. contour_integral (linepath a b) (F w))" proof - have *: "\d>0. \x'\U. dist x' w < d \ dist (contour_integral (linepath a b) (F x')) (contour_integral (linepath a b) (F w)) \ \" if "w \ U" "0 < \" "a \ b" for w \ proof - obtain \ where "\>0" and \: "cball w \ \ U" using open_contains_cball \open U\ \w \ U\ by force let ?TZ = "cball w \ \ closed_segment a b" have "uniformly_continuous_on ?TZ (\(x,y). F x y)" proof (rule compact_uniformly_continuous) show "continuous_on ?TZ (\(x,y). F x y)" by (rule continuous_on_subset[OF cond_uu]) (use SigmaE \ abu in blast) show "compact ?TZ" by (simp add: compact_Times) qed then obtain \ where "\>0" and \: "\x x'. \x\?TZ; x'\?TZ; dist x' x < \\ \ dist ((\(x,y). F x y) x') ((\(x,y). F x y) x) < \/norm(b - a)" using \0 < \\ \a \ b\ by (auto elim: uniformly_continuous_onE [where e = "\/norm(b - a)"]) have \: "\norm (w - x1) \ \; x2 \ closed_segment a b; norm (w - x1') \ \; x2' \ closed_segment a b; norm ((x1', x2') - (x1, x2)) < \\ \ norm (F x1' x2' - F x1 x2) \ \ / cmod (b - a)" for x1 x2 x1' x2' using \ [of "(x1,x2)" "(x1',x2')"] by (force simp: dist_norm) have le_ee: "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \" if "x' \ U" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' proof - have "(\x. F x' x - F w x) contour_integrable_on linepath a b" by (simp add: \w \ U\ cont_dw contour_integrable_diff that) then have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \]) using \0 < \\ \0 < \\ that by (auto simp: norm_minus_commute) also have "\ = \" using \a \ b\ by simp finally show ?thesis . qed show ?thesis apply (rule_tac x="min \ \" in exI) using \0 < \\ \0 < \\ by (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \w \ U\ intro: le_ee) qed show ?thesis proof (cases "a=b") case False show ?thesis by (rule continuous_onI) (use False in \auto intro: *\) qed auto qed text\This version has \<^term>\polynomial_function \\ as an additional assumption.\ lemma Cauchy_integral_formula_global_weak: assumes "open U" and holf: "f holomorphic_on U" and z: "z \ U" and \: "polynomial_function \" and pasz: "path_image \ \ U - {z}" and loop: "pathfinish \ = pathstart \" and zero: "\w. w \ U \ winding_number \ w = 0" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - obtain \' where pf\': "polynomial_function \'" and \': "\x. (\ has_vector_derivative (\' x)) (at x)" using has_vector_derivative_polynomial_function [OF \] by blast then have "bounded(path_image \')" by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" using bounded_pos by force 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 define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" have "path \" "valid_path \" using \ by (auto simp: path_polynomial_function valid_path_polynomial_function) then have ov: "open v" by (simp add: v_def open_winding_number_levelsets loop) have uv_Un: "U \ v = UNIV" using pasz zero by (auto simp: v_def) have conf: "continuous_on U f" by (metis holf holomorphic_on_imp_continuous_on) have hol_d: "(d y) holomorphic_on U" if "y \ U" for y proof - have *: "(\c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on U" by (simp add: holf pole_lemma_open \open U\) then have "isCont (\x. if x = y then deriv f y else (f x - f y) / (x - y)) y" using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that \open U\ by fastforce then have "continuous_on U (d y)" using "*" d_def holomorphic_on_imp_continuous_on by auto moreover have "d y holomorphic_on U - {y}" proof - have "(\w. if w = y then deriv f y else (f w - f y) / (w - y)) field_differentiable at w" if "w \ U - {y}" for w proof (rule field_differentiable_transform_within) show "(\w. (f w - f y) / (w - y)) field_differentiable at w" using that \open U\ holf by (auto intro!: holomorphic_on_imp_differentiable_at derivative_intros) show "dist w y > 0" using that by auto qed (auto simp: dist_commute) then show ?thesis unfolding field_differentiable_def by (simp add: d_def holomorphic_on_open \open U\ open_delete) qed ultimately show ?thesis by (rule no_isolated_singularity) (auto simp: \open U\) qed have cint_fxy: "(\x. (f x - f y) / (x - y)) contour_integrable_on \" if "y \ path_image \" for y proof (rule contour_integrable_holomorphic_simple [where S = "U-{y}"]) show "(\x. (f x - f y) / (x - y)) holomorphic_on U - {y}" by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) show "path_image \ \ U - {y}" using pasz that by blast qed (auto simp: \open U\ open_delete \valid_path \\) define h where "h z = (if z \ U then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z have U: "((d z) has_contour_integral h z) \" if "z \ U" for z proof - have "d z holomorphic_on U" by (simp add: hol_d that) with that show ?thesis by (metis Diff_subset \valid_path \\ \open U\ contour_integrable_holomorphic_simple h_def has_contour_integral_integral pasz subset_trans) qed have V: "((\w. f w / (w - z)) has_contour_integral h z) \" if z: "z \ v" for z proof - have 0: "0 = (f z) * 2 * of_real (2 * pi) * \ * winding_number \ z" using v_def z by auto then have "((\x. 1 / (x - z)) has_contour_integral 0) \" using z v_def has_contour_integral_winding_number [OF \valid_path \\] by fastforce then have "((\x. f z * (1 / (x - z))) has_contour_integral 0) \" using has_contour_integral_lmul by fastforce then have "((\x. f z / (x - z)) has_contour_integral 0) \" by (simp add: field_split_simps) moreover have "((\x. (f x - f z) / (x - z)) has_contour_integral contour_integral \ (d z)) \" using z apply (simp add: v_def) apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) done ultimately have *: "((\x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \ (d z))) \" by (rule has_contour_integral_add) have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (d z)) \" if "z \ U" using * by (auto simp: divide_simps has_contour_integral_eq) moreover have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (\w. f w / (w - z))) \" if "z \ U" proof (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where S=U]]) show "(\w. f w / (w - z)) holomorphic_on U" by (rule holomorphic_intros assms | use that in force)+ qed (use \open U\ pasz \valid_path \\ in auto) ultimately show ?thesis using z by (simp add: h_def) qed have znot: "z \ path_image \" using pasz by blast obtain d0 where "d0>0" and d0: "\x y. x \ path_image \ \ y \ - U \ d0 \ dist x y" using separate_compact_closed [of "path_image \" "-U"] pasz \open U\ \path \\ compact_path_image by blast obtain dd where "0 < dd" and dd: "{y + k | y k. y \ path_image \ \ k \ ball 0 dd} \ U" proof show "0 < d0 / 2" using \0 < d0\ by auto qed (use \0 < d0\ d0 in \force simp: dist_norm\) define T where "T \ {y + k |y k. y \ path_image \ \ k \ cball 0 (dd / 2)}" have "\x x'. \x \ path_image \; dist x x' * 2 < dd\ \ \y k. x' = y + k \ y \ path_image \ \ dist 0 k * 2 \ dd" apply (rule_tac x=x in exI) apply (rule_tac x="x'-x" in exI) apply (force simp: dist_norm) done then have subt: "path_image \ \ interior T" using \0 < dd\ apply (clarsimp simp add: mem_interior T_def) apply (rule_tac x="dd/2" in exI, auto) done have "compact T" unfolding T_def by (fastforce simp add: \valid_path \\ compact_valid_path_image intro!: compact_sums) have T: "T \ U" unfolding T_def using \0 < dd\ dd by fastforce obtain L where "L>0" and L: "\f B. \f holomorphic_on interior T; \z. z\interior T \ cmod (f z) \ B\ \ cmod (contour_integral \ f) \ L * B" using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] by blast have "bounded(f ` T)" by (meson \compact T\ compact_continuous_image compact_imp_bounded conf continuous_on_subset T) then obtain D where "D>0" and D: "\x. x \ T \ norm (f x) \ D" by (auto simp: bounded_pos) obtain C where "C>0" and C: "\x. x \ T \ norm x \ C" using \compact T\ bounded_pos compact_imp_bounded by force have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y proof - have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp with le have ybig: "norm y > C" by force with C have "y \ T" by force then have ynot: "y \ path_image \" using subt interior_subset by blast have [simp]: "winding_number \ y = 0" proof (rule winding_number_zero_outside) show "path_image \ \ cball 0 C" by (meson C interior_subset mem_cball_0 subset_eq subt) qed (use ybig loop \path \\ in auto) have [simp]: "h y = contour_integral \ (\w. f w/(w - y))" by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V) have holint: "(\w. f w / (w - y)) holomorphic_on interior T" proof (intro holomorphic_intros) show "f holomorphic_on interior T" using holf holomorphic_on_subset interior_subset T by blast qed (use \y \ T\ interior_subset in auto) have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior T" for z proof - have "D * L / e + cmod z \ cmod y" using le C [of z] z using interior_subset by force then have DL2: "D * L / e \ cmod (z - y)" using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) also have "\ \ D * (e / L / D)" proof (rule mult_mono) show "cmod (f z) \ D" using D interior_subset z by blast show "inverse (cmod (z - y)) \ e / L / D" "D \ 0" using \L>0\ \e>0\ \D>0\ DL2 by (auto simp: norm_divide field_split_simps) qed auto finally show ?thesis . qed have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" by (simp add: dist_norm) also have "\ \ L * (D * (e / L / D))" by (rule L [OF holint leD]) also have "\ = e" using \L>0\ \0 < D\ by auto finally show ?thesis . qed then have "(h \ 0) at_infinity" by (meson Lim_at_infinityI) moreover have "h holomorphic_on UNIV" proof - have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" if "x \ U" "z \ U" "x \ z" for x z using that conf apply (simp add: split_def continuous_on_eq_continuous_at \open U\) apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ done have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" by (rule continuous_intros)+ have open_uu_Id: "open (U \ U - Id)" proof (rule open_Diff) show "open (U \ U)" by (simp add: open_Times \open U\) show "closed (Id :: complex rel)" using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] by (auto simp: Id_fstsnd_eq algebra_simps) qed have con_derf: "continuous (at z) (deriv f)" if "z \ U" for z proof (rule continuous_on_interior) show "continuous_on U (deriv f)" by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open U\) qed (simp add: interior_open that \open U\) have tendsto_f': "((\(x,y). if y = x then deriv f (x) else (f (y) - f (x)) / (y - x)) \ deriv f x) (at (x, x) within U \ U)" if "x \ U" for x proof (rule Lim_withinI) fix e::real assume "0 < e" obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" using \0 < e\ continuous_within_E [OF con_derf [OF \x \ U\]] by (metis UNIV_I dist_norm) obtain k2 where "k2>0" and k2: "ball x k2 \ U" by (blast intro: openE [OF \open U\] \x \ U\) have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" for x' z' proof - have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w using segment_furthest_le [of w x' z' x] by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) have f_has_der: "\x. x \ U \ (f has_field_derivative deriv f x) (at x within U)" by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \open U\) have "closed_segment x' z' \ U" by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" by (rule has_contour_integral_div) have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] \e > 0\ \z' \ x'\ apply (auto simp: norm_divide divide_simps derf_le) done also have "\ \ e" using \0 < e\ by simp finally show ?thesis . qed show "\d>0. \xa\U \ U. 0 < dist xa (x, x) \ dist xa (x, x) < d \ dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" apply (rule_tac x="min k1 k2" in exI) using \k1>0\ \k2>0\ \e>0\ by (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) qed have con_pa_f: "continuous_on (path_image \) f" by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt T) have le_B: "\T. T \ {0..1} \ cmod (vector_derivative \ (at T)) \ B" using \' B by (simp add: path_image_def vector_derivative_at rev_image_eqI) have f_has_cint: "\w. w \ v - path_image \ \ ((\u. f u / (u - w) ^ 1) has_contour_integral h w) \" by (simp add: V) have cond_uu: "continuous_on (U \ U) (\(x,y). d x y)" apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') apply (simp add: tendsto_within_open_NO_MATCH open_Times \open U\, clarify) apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) using con_ff apply (auto simp: continuous_within) done have hol_dw: "(\z. d z w) holomorphic_on U" if "w \ U" for w proof - have "continuous_on U ((\(x,y). d x y) \ (\z. (w,z)))" by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ then have *: "continuous_on U (\z. if w = z then deriv f z else (f w - f z) / (w - z))" by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) have **: "(\z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" if "x \ U" "x \ w" for x proof (rule_tac f = "\x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) show "(\x. (f w - f x) / (w - x)) field_differentiable at x" using that \open U\ by (intro derivative_intros holomorphic_on_imp_differentiable_at [OF holf]; force) qed (use that \open U\ in \auto simp: dist_commute\) show ?thesis unfolding d_def proof (rule no_isolated_singularity [OF * _ \open U\]) show "(\z. if w = z then deriv f z else (f w - f z) / (w - z)) holomorphic_on U - {w}" by (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff \open U\ **) qed auto qed { fix a b assume abu: "closed_segment a b \ U" have cont_cint_d: "continuous_on U (\w. contour_integral (linepath a b) (\z. d z w))" proof (rule contour_integral_continuous_on_linepath_2D [OF \open U\ _ _ abu]) show "\w. w \ U \ (\z. d z w) contour_integrable_on (linepath a b)" by (metis abu hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) show "continuous_on (U \ U) (\(x, y). d y x)" by (auto intro: continuous_on_swap_args cond_uu) qed have cont_cint_d\: "continuous_on {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) \ \)" proof (rule continuous_on_compose) show "continuous_on {0..1} \" using \path \\ path_def by blast show "continuous_on (\ ` {0..1}) (\w. contour_integral (linepath a b) (\z. d z w))" using pasz unfolding path_image_def by (auto intro!: continuous_on_subset [OF cont_cint_d]) qed have "continuous_on {0..1} (\x. vector_derivative \ (at x))" using pf\' by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \']) then have cint_cint: "(\w. contour_integral (linepath a b) (\z. d z w)) contour_integrable_on \" apply (simp add: contour_integrable_on) apply (rule integrable_continuous_real) by (rule continuous_on_mult [OF cont_cint_d\ [unfolded o_def]]) have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\z. contour_integral \ (d z))" using abu by (force simp: h_def intro: contour_integral_eq) also have "\ = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" proof (rule contour_integral_swap) show "continuous_on (path_image (linepath a b) \ path_image \) (\(y1, y2). d y1 y2)" using abu pasz by (auto intro: continuous_on_subset [OF cond_uu]) show "continuous_on {0..1} (\t. vector_derivative (linepath a b) (at t))" by (auto intro!: continuous_intros) show "continuous_on {0..1} (\t. vector_derivative \ (at t))" by (metis \' continuous_on_eq path_def path_polynomial_function pf\' vector_derivative_at) qed (use \valid_path \\ in auto) finally have cint_h_eq: "contour_integral (linepath a b) h = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . note cint_cint cint_h_eq } note cint_h = this have conthu: "continuous_on U h" proof (simp add: continuous_on_sequentially, clarify) fix a x assume x: "x \ U" and au: "\n. a n \ U" and ax: "a \ x" then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" by (meson U contour_integrable_on_def eventuallyI) obtain dd where "dd>0" and dd: "cball x dd \ U" using open_contains_cball \open U\ x by force have A2: "uniform_limit (path_image \) (\n. d (a n)) (d x) sequentially" unfolding uniform_limit_iff dist_norm proof clarify fix ee::real assume "0 < ee" show "\\<^sub>F n in sequentially. \\\path_image \. cmod (d (a n) \ - d x \) < ee" proof - let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" proof (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) show "compact {(w, z) |w z. w \ cball x dd \ z \ path_image \}" using \valid_path \\ by (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) qed (use dd pasz in auto) then obtain kk where "kk>0" and kk: "\x x'. \x \ ?ddpa; x' \ ?ddpa; dist x' x < kk\ \ dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" by (rule uniformly_continuous_onE [where e = ee]) (use \0 < ee\ in auto) have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" for w z using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm) obtain no where "\n\no. dist (a n) x < min dd kk" using ax unfolding lim_sequentially by (meson \0 < dd\ \0 < kk\ min_less_iff_conj) then show ?thesis using \dd > 0\ \kk > 0\ by (fastforce simp: eventually_sequentially kk dist_norm) qed qed have "(\n. contour_integral \ (d (a n))) \ contour_integral \ (d x)" by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \valid_path \\) then have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" by (simp add: h_def x) then show "(h \ a) \ h x" by (simp add: h_def x au o_def) qed show ?thesis proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) fix z0 consider "z0 \ v" | "z0 \ U" using uv_Un by blast then show "h field_differentiable at z0" proof cases assume "z0 \ v" then show ?thesis using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ by (auto simp: field_differentiable_def v_def) next assume "z0 \ U" then obtain e where "e>0" and e: "ball z0 e \ U" by (blast intro: openE [OF \open U\]) have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c proof - have *: "\x1 x2 z. z \ U \ closed_segment x1 x2 \ U \ (\w. d w z) contour_integrable_on linepath x1 x2" using hol_dw holomorphic_on_imp_continuous_on \open U\ by (auto intro!: contour_integrable_holomorphic_simple) have abc: "closed_segment a b \ U" "closed_segment b c \ U" "closed_segment c a \ U" using that e segments_subset_convex_hull by fastforce+ have eq0: "\w. w \ U \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" proof (rule contour_integral_unique [OF Cauchy_theorem_triangle]) show "\w. w \ U \ (\z. d z w) holomorphic_on convex hull {a, b, c}" using e abc_subset by (auto intro: holomorphic_on_subset [OF hol_dw]) qed have "contour_integral \ (\x. contour_integral (linepath a b) (\z. d z x) + (contour_integral (linepath b c) (\z. d z x) + contour_integral (linepath c a) (\z. d z x))) = 0" apply (rule contour_integral_eq_0) using abc pasz U apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ done then show ?thesis by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) qed show ?thesis using e \e > 0\ by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic Morera_triangle continuous_on_subset [OF conthu] *) qed qed qed ultimately have [simp]: "h z = 0" for z by (meson Liouville_weak) have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" by (metis mult.commute has_contour_integral_lmul) then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" by (simp add: field_split_simps) moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) show ?thesis using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) qed theorem Cauchy_integral_formula_global: assumes S: "open S" and holf: "f holomorphic_on S" and z: "z \ S" and vpg: "valid_path \" and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" and zero: "\w. w \ S \ winding_number \ w = 0" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - have "path \" using vpg by (blast intro: valid_path_imp_path) have hols: "(\w. f w / (w - z)) holomorphic_on S - {z}" "(\w. 1 / (w - z)) holomorphic_on S - {z}" by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete S vpg pasz) obtain d where "d>0" and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; pathstart h = pathstart g \ pathfinish h = pathfinish g\ \ path_image h \ S - {z} \ (\f. f holomorphic_on S - {z} \ contour_integral h f = contour_integral g f)" using contour_integral_nearby_ends [OF _ \path \\ pasz] S by (simp add: open_Diff) metis obtain p where polyp: "polynomial_function p" and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" using path_approx_polynomial_function [OF \path \\ \d > 0\] by metis then have ploop: "pathfinish p = pathstart p" using loop by auto have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast have [simp]: "z \ path_image \" using pasz by blast have paps: "path_image p \ S - {z}" and cint_eq: "(\f. f holomorphic_on S - {z} \ contour_integral p f = contour_integral \ f)" using pf ps led d [OF vpg vpp] \d > 0\ by auto have wn_eq: "winding_number p z = winding_number \ z" using vpp paps by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) have "winding_number p w = winding_number \ w" if "w \ S" for w proof - have hol: "(\v. 1 / (v - w)) holomorphic_on S - {z}" using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) have "w \ path_image p" "w \ path_image \" using paps pasz that by auto then show ?thesis using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) qed then have wn0: "\w. w \ S \ winding_number p w = 0" by (simp add: zero) show ?thesis using Cauchy_integral_formula_global_weak [OF S holf z polyp paps ploop wn0] hols by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) qed theorem Cauchy_theorem_global: assumes S: "open S" and holf: "f holomorphic_on S" and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" and pas: "path_image \ \ S" and zero: "\w. w \ S \ winding_number \ w = 0" shows "(f has_contour_integral 0) \" proof - obtain z where "z \ S" and znot: "z \ path_image \" proof - have "compact (path_image \)" using compact_valid_path_image vpg by blast then have "path_image \ \ S" by (metis (no_types) compact_open path_image_nonempty S) with pas show ?thesis by (blast intro: that) qed then have pasz: "path_image \ \ S - {z}" using pas by blast have hol: "(\w. (w - z) * f w) holomorphic_on S" by (rule holomorphic_intros holf)+ show ?thesis using Cauchy_integral_formula_global [OF S hol \z \ S\ vpg pasz loop zero] by (auto simp: znot elim!: has_contour_integral_eq) qed corollary Cauchy_theorem_global_outside: assumes "open S" "f holomorphic_on S" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ S" "\w. w \ S \ w \ outside(path_image \)" shows "(f has_contour_integral 0) \" by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) lemma simply_connected_imp_winding_number_zero: assumes "simply_connected S" "path g" "path_image g \ S" "pathfinish g = pathstart g" "z \ S" shows "winding_number g z = 0" proof - have hom: "homotopic_loops S g (linepath (pathstart g) (pathstart g))" by (meson assms homotopic_paths_imp_homotopic_loops pathfinish_linepath simply_connected_eq_contractible_path) then have "homotopic_paths (- {z}) g (linepath (pathstart g) (pathstart g))" by (meson \z \ S\ homotopic_loops_imp_homotopic_paths_null homotopic_paths_subset subset_Compl_singleton) then have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z" by (rule winding_number_homotopic_paths) also have "\ = 0" using assms by (force intro: winding_number_trivial) finally show ?thesis . qed lemma Cauchy_theorem_simply_connected: assumes "open S" "simply_connected S" "f holomorphic_on S" "valid_path g" "path_image g \ S" "pathfinish g = pathstart g" shows "(f has_contour_integral 0) g" using assms apply (simp add: simply_connected_eq_contractible_path) apply (auto intro!: Cauchy_theorem_null_homotopic [where a = "pathstart g"] homotopic_paths_imp_homotopic_loops) using valid_path_imp_path by blast proposition\<^marker>\tag unimportant\ holomorphic_logarithm_exists: assumes A: "convex A" "open A" and f: "f holomorphic_on A" "\x. x \ A \ f x \ 0" and z0: "z0 \ A" obtains g where "g holomorphic_on A" and "\x. x \ A \ exp (g x) = f x" proof - note f' = holomorphic_derivI [OF f(1) A(2)] obtain g where g: "\x. x \ A \ (g has_field_derivative deriv f x / f x) (at x)" proof (rule holomorphic_convex_primitive' [OF A]) show "(\x. deriv f x / f x) holomorphic_on A" by (intro holomorphic_intros f A) qed (auto simp: A at_within_open[of _ A]) define h where "h = (\x. -g z0 + ln (f z0) + g x)" from g and A have g_holo: "g holomorphic_on A" by (auto simp: holomorphic_on_def at_within_open[of _ A] field_differentiable_def) hence h_holo: "h holomorphic_on A" by (auto simp: h_def intro!: holomorphic_intros) have "\c. \x\A. f x / exp (h x) - 1 = c" proof (rule has_field_derivative_zero_constant, goal_cases) case (2 x) note [simp] = at_within_open[OF _ \open A\] from 2 and z0 and f show ?case by (auto simp: h_def exp_diff field_simps intro!: derivative_eq_intros g f') qed fact+ then obtain c where c: "\x. x \ A \ f x / exp (h x) - 1 = c" by blast from c[OF z0] and z0 and f have "c = 0" by (simp add: h_def) with c have "\x. x \ A \ exp (h x) = f x" by simp from that[OF h_holo this] show ?thesis . qed (* FIXME mv to Cauchy_Integral_Theorem.thy *) subsection\Cauchy's inequality and more versions of Liouville\ lemma Cauchy_higher_deriv_bound: assumes holf: "f holomorphic_on (ball z r)" and contf: "continuous_on (cball z r) f" and fin : "\w. w \ ball z r \ f w \ ball y B0" and "0 < r" and "0 < n" shows "norm ((deriv ^^ n) f z) \ (fact n) * B0 / r^n" proof - have "0 < B0" using \0 < r\ fin [of z] by (metis ball_eq_empty ex_in_conv fin not_less) have le_B0: "cmod (f w - y) \ B0" if "cmod (w - z) \ r" for w proof (rule continuous_on_closure_norm_le [of "ball z r" "\w. f w - y"], use \0 < r\ in simp_all) show "continuous_on (cball z r) (\w. f w - y)" by (intro continuous_intros contf) show "dist z w \ r" by (simp add: dist_commute dist_norm that) qed (use fin in \auto simp: dist_norm less_eq_real_def norm_minus_commute\) have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w) z - (deriv ^^ n) (\w. y) z" using \0 < n\ by simp also have "... = (deriv ^^ n) (\w. f w - y) z" by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \0 < r\) finally have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w - y) z" . have contf': "continuous_on (cball z r) (\u. f u - y)" by (rule contf continuous_intros)+ have holf': "(\u. (f u - y)) holomorphic_on (ball z r)" by (simp add: holf holomorphic_on_diff) define a where "a = (2 * pi)/(fact n)" have "0 < a" by (simp add: a_def) have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" using \0 < r\ by (simp add: a_def field_split_simps) have der_dif: "(deriv ^^ n) (\w. f w - y) z = (deriv ^^ n) f z" using \0 < r\ \0 < n\ by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) have "norm ((2 * of_real pi * \)/(fact n) * (deriv ^^ n) (\w. f w - y) z) \ (B0/r^(Suc n)) * (2 * pi * r)" apply (rule has_contour_integral_bound_circlepath [of "(\u. (f u - y)/(u - z)^(Suc n))" _ z]) using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] using \0 < B0\ \0 < r\ apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) done then show ?thesis using \0 < r\ by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) qed lemma Cauchy_inequality: assumes holf: "f holomorphic_on (ball \ r)" and contf: "continuous_on (cball \ r) f" and "0 < r" and nof: "\x. norm(\-x) = r \ norm(f x) \ B" shows "norm ((deriv ^^ n) f \) \ (fact n) * B / r^n" proof - obtain x where "norm (\-x) = r" by (metis abs_of_nonneg add_diff_cancel_left' \0 < r\ diff_add_cancel dual_order.strict_implies_order norm_of_real) then have "0 \ B" by (metis nof norm_not_less_zero not_le order_trans) have "\ \ ball \ r" using \0 < r\ by simp then have "((\u. f u / (u-\) ^ Suc n) has_contour_integral (2 * pi) * \ / fact n * (deriv ^^ n) f \) (circlepath \ r)" by (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) have "norm ((2 * pi * \)/(fact n) * (deriv ^^ n) f \) \ (B / r^(Suc n)) * (2 * pi * r)" proof (rule has_contour_integral_bound_circlepath) have "\ \ ball \ r" using \0 < r\ by simp then show "((\u. f u / (u-\) ^ Suc n) has_contour_integral (2 * pi) * \ / fact n * (deriv ^^ n) f \) (circlepath \ r)" by (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) show "\x. cmod (x-\) = r \ cmod (f x / (x-\) ^ Suc n) \ B / r ^ Suc n" using \0 \ B\ \0 < r\ by (simp add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc) qed (use \0 \ B\ \0 < r\ in auto) then show ?thesis using \0 < r\ by (simp add: norm_divide norm_mult field_simps) qed lemma Liouville_polynomial: assumes holf: "f holomorphic_on UNIV" and nof: "\z. A \ norm z \ norm(f z) \ B * norm z ^ n" shows "f \ = (\k\n. (deriv^^k) f 0 / fact k * \ ^ k)" proof (cases rule: le_less_linear [THEN disjE]) assume "B \ 0" then have "\z. A \ norm z \ norm(f z) = 0" by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le) then have f0: "(f \ 0) at_infinity" using Lim_at_infinity by force then have [simp]: "f = (\w. 0)" using Liouville_weak [OF holf, of 0] by (simp add: eventually_at_infinity f0) meson show ?thesis by simp next assume "0 < B" have "((\k. (deriv ^^ k) f 0 / (fact k) * (\ - 0)^k) sums f \)" proof (rule holomorphic_power_series [where r = "norm \ + 1"]) show "f holomorphic_on ball 0 (cmod \ + 1)" "\ \ ball 0 (cmod \ + 1)" using holf holomorphic_on_subset by auto qed then have sumsf: "((\k. (deriv ^^ k) f 0 / (fact k) * \^k) sums f \)" by simp have "(deriv ^^ k) f 0 / fact k * \ ^ k = 0" if "k>n" for k proof (cases "(deriv ^^ k) f 0 = 0") case True then show ?thesis by simp next case False define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" have "1 \ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" using \0 < B\ by simp then have wge1: "1 \ norm w" by (metis norm_of_real w_def) then have "w \ 0" by auto have kB: "0 < fact k * B" using \0 < B\ by simp then have "0 \ fact k * B / cmod ((deriv ^^ k) f 0)" by simp then have wgeA: "A \ cmod w" by (simp only: w_def norm_of_real) have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" using \0 < B\ by simp then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" by (metis norm_of_real w_def) then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" using False by (simp add: field_split_simps mult.commute split: if_split_asm) also have "... \ fact k * (B * norm w ^ n) / norm w ^ k" proof (rule Cauchy_inequality) show "f holomorphic_on ball 0 (cmod w)" using holf holomorphic_on_subset by force show "continuous_on (cball 0 (cmod w)) f" using holf holomorphic_on_imp_continuous_on holomorphic_on_subset by blast show "\x. cmod (0 - x) = cmod w \ cmod (f x) \ B * cmod w ^ n" by (metis nof wgeA dist_0_norm dist_norm) qed (use \w \ 0\ in auto) also have "... = fact k * (B * 1 / cmod w ^ (k-n))" proof - have "cmod w ^ n / cmod w ^ k = 1 / cmod w ^ (k - n)" using \k>n\ \w \ 0\ \0 < B\ by (simp add: field_split_simps semiring_normalization_rules) then show ?thesis by (metis times_divide_eq_right) qed also have "... = fact k * B / cmod w ^ (k-n)" by simp finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . then have "1 / cmod w < 1 / cmod w ^ (k - n)" by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) then have "cmod w ^ (k - n) < cmod w" by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) with self_le_power [OF wge1] have False by (meson diff_is_0_eq not_gr0 not_le that) then show ?thesis by blast qed then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \ ^ (k + Suc n) = 0" for k using not_less_eq by blast then have "(\i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \ ^ (i + Suc n)) sums 0" by (rule sums_0) with sums_split_initial_segment [OF sumsf, where n = "Suc n"] show ?thesis using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce qed text\Every bounded entire function is a constant function.\ theorem Liouville_theorem: assumes holf: "f holomorphic_on UNIV" and bf: "bounded (range f)" - obtains c where "\z. f z = c" + shows "f constant_on UNIV" proof - obtain B where "\z. cmod (f z) \ B" by (meson bf bounded_pos rangeI) then show ?thesis - using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast + using Liouville_polynomial [OF holf, of 0 B 0, simplified] + by (meson constant_on_def) qed text\A holomorphic function f has only isolated zeros unless f is 0.\ lemma powser_0_nonzero: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes r: "0 < r" and sm: "\x. norm (x-\) < r \ (\n. a n * (x-\) ^ n) sums (f x)" and [simp]: "f \ = 0" and m0: "a m \ 0" and "m>0" obtains s where "0 < s" and "\z. z \ cball \ s - {\} \ f z \ 0" proof - have "r \ conv_radius a" using sm sums_summable by (auto simp: le_conv_radius_iff [where \=\]) obtain m where am: "a m \ 0" and az [simp]: "(\n. n a n = 0)" proof show "a (LEAST n. a n \ 0) \ 0" by (metis (mono_tags, lifting) m0 LeastI) qed (fastforce dest!: not_less_Least) define b where "b i = a (i+m) / a m" for i define g where "g x = suminf (\i. b i * (x-\) ^ i)" for x have [simp]: "b 0 = 1" by (simp add: am b_def) { fix x::'a assume "norm (x-\) < r" then have "(\n. (a m * (x-\)^m) * (b n * (x-\)^n)) sums (f x)" using am az sm sums_zero_iff_shift [of m "(\n. a n * (x-\) ^ n)" "f x"] by (simp add: b_def monoid_mult_class.power_add algebra_simps) then have "x \ \ \ (\n. b n * (x-\)^n) sums (f x / (a m * (x-\)^m))" using am by (simp add: sums_mult_D) } note bsums = this then have "norm (x-\) < r \ summable (\n. b n * (x-\)^n)" for x using sums_summable by (cases "x=\") auto then have "r \ conv_radius b" by (simp add: le_conv_radius_iff [where \=\]) then have "r/2 < conv_radius b" using not_le order_trans r by fastforce then have "continuous_on (cball \ (r/2)) g" using powser_continuous_suminf [of "r/2" b \] by (simp add: g_def) then obtain s where "s>0" "\x. \norm (x-\) \ s; norm (x-\) \ r/2\ \ dist (g x) (g \) < 1/2" proof (rule continuous_onE) show "\ \ cball \ (r / 2)" "1/2 > (0::real)" using r by auto qed (auto simp: dist_commute dist_norm) moreover have "g \ = 1" by (simp add: g_def) ultimately have gnz: "\x. \norm (x-\) \ s; norm (x-\) \ r/2\ \ (g x) \ 0" by fastforce have "f x \ 0" if "x \ \" "norm (x-\) \ s" "norm (x-\) \ r/2" for x using bsums [of x] that gnz [of x] r sums_iff unfolding g_def by fastforce then show ?thesis apply (rule_tac s="min s (r/2)" in that) using \0 < r\ \0 < s\ by (auto simp: dist_commute dist_norm) qed subsection \Complex functions and power series\ text \ The following defines the power series expansion of a complex function at a given point (assuming that it is analytic at that point). \ definition\<^marker>\tag important\ fps_expansion :: "(complex \ complex) \ complex \ complex fps" where "fps_expansion f z0 = Abs_fps (\n. (deriv ^^ n) f z0 / fact n)" lemma fixes r :: ereal assumes "f holomorphic_on eball z0 r" shows conv_radius_fps_expansion: "fps_conv_radius (fps_expansion f z0) \ r" and eval_fps_expansion: "\z. z \ eball z0 r \ eval_fps (fps_expansion f z0) (z - z0) = f z" and eval_fps_expansion': "\z. norm z < r \ eval_fps (fps_expansion f z0) z = f (z0 + z)" proof - have "(\n. fps_nth (fps_expansion f z0) n * (z - z0) ^ n) sums f z" if "z \ ball z0 r'" "ereal r' < r" for z r' proof - from that(2) have "ereal r' \ r" by simp from assms(1) and this have "f holomorphic_on ball z0 r'" by (rule holomorphic_on_subset[OF _ ball_eball_mono]) from holomorphic_power_series [OF this that(1)] show ?thesis by (simp add: fps_expansion_def) qed hence *: "(\n. fps_nth (fps_expansion f z0) n * (z - z0) ^ n) sums f z" if "z \ eball z0 r" for z using that by (subst (asm) eball_conv_UNION_balls) blast show "fps_conv_radius (fps_expansion f z0) \ r" unfolding fps_conv_radius_def proof (rule conv_radius_geI_ex) fix r' :: real assume r': "r' > 0" "ereal r' < r" thus "\z. norm z = r' \ summable (\n. fps_nth (fps_expansion f z0) n * z ^ n)" using *[of "z0 + of_real r'"] by (intro exI[of _ "of_real r'"]) (auto simp: summable_def dist_norm) qed show "eval_fps (fps_expansion f z0) (z - z0) = f z" if "z \ eball z0 r" for z using *[OF that] by (simp add: eval_fps_def sums_iff) show "eval_fps (fps_expansion f z0) z = f (z0 + z)" if "ereal (norm z) < r" for z using *[of "z0 + z"] and that by (simp add: eval_fps_def sums_iff dist_norm) qed text \ We can now show several more facts about power series expansions (at least in the complex case) with relative ease that would have been trickier without complex analysis. \ lemma fixes f :: "complex fps" and r :: ereal assumes "\z. ereal (norm z) < r \ eval_fps f z \ 0" shows fps_conv_radius_inverse: "fps_conv_radius (inverse f) \ min r (fps_conv_radius f)" and eval_fps_inverse: "\z. ereal (norm z) < fps_conv_radius f \ ereal (norm z) < r \ eval_fps (inverse f) z = inverse (eval_fps f z)" proof - define R where "R = min (fps_conv_radius f) r" have *: "fps_conv_radius (inverse f) \ min r (fps_conv_radius f) \ (\z\eball 0 (min (fps_conv_radius f) r). eval_fps (inverse f) z = inverse (eval_fps f z))" proof (cases "min r (fps_conv_radius f) > 0") case True define f' where "f' = fps_expansion (\z. inverse (eval_fps f z)) 0" have holo: "(\z. inverse (eval_fps f z)) holomorphic_on eball 0 (min r (fps_conv_radius f))" using assms by (intro holomorphic_intros) auto from holo have radius: "fps_conv_radius f' \ min r (fps_conv_radius f)" unfolding f'_def by (rule conv_radius_fps_expansion) have eval_f': "eval_fps f' z = inverse (eval_fps f z)" if "norm z < fps_conv_radius f" "norm z < r" for z using that unfolding f'_def by (subst eval_fps_expansion'[OF holo]) auto have "f * f' = 1" proof (rule eval_fps_eqD) from radius and True have "0 < min (fps_conv_radius f) (fps_conv_radius f')" by (auto simp: min_def split: if_splits) also have "\ \ fps_conv_radius (f * f')" by (rule fps_conv_radius_mult) finally show "\ > 0" . next from True have "R > 0" by (auto simp: R_def) hence "eventually (\z. z \ eball 0 R) (nhds 0)" by (intro eventually_nhds_in_open) (auto simp: zero_ereal_def) thus "eventually (\z. eval_fps (f * f') z = eval_fps 1 z) (nhds 0)" proof eventually_elim case (elim z) hence "eval_fps (f * f') z = eval_fps f z * eval_fps f' z" using radius by (intro eval_fps_mult) (auto simp: R_def min_def split: if_splits intro: less_trans) also have "eval_fps f' z = inverse (eval_fps f z)" using elim by (intro eval_f') (auto simp: R_def) also from elim have "eval_fps f z \ 0" by (intro assms) (auto simp: R_def) hence "eval_fps f z * inverse (eval_fps f z) = eval_fps 1 z" by simp finally show "eval_fps (f * f') z = eval_fps 1 z" . qed qed simp_all hence "f' = inverse f" by (intro fps_inverse_unique [symmetric]) (simp_all add: mult_ac) with eval_f' and radius show ?thesis by simp next case False hence *: "eball 0 R = {}" by (intro eball_empty) (auto simp: R_def min_def split: if_splits) show ?thesis proof safe from False have "min r (fps_conv_radius f) \ 0" by (simp add: min_def) also have "0 \ fps_conv_radius (inverse f)" by (simp add: fps_conv_radius_def conv_radius_nonneg) finally show "min r (fps_conv_radius f) \ \" . qed (unfold * [unfolded R_def], auto) qed from * show "fps_conv_radius (inverse f) \ min r (fps_conv_radius f)" by blast from * show "eval_fps (inverse f) z = inverse (eval_fps f z)" if "ereal (norm z) < fps_conv_radius f" "ereal (norm z) < r" for z using that by auto qed lemma fixes f g :: "complex fps" and r :: ereal defines "R \ Min {r, fps_conv_radius f, fps_conv_radius g}" assumes "fps_conv_radius f > 0" "fps_conv_radius g > 0" "r > 0" assumes nz: "\z. z \ eball 0 r \ eval_fps g z \ 0" shows fps_conv_radius_divide': "fps_conv_radius (f / g) \ R" and eval_fps_divide': "ereal (norm z) < R \ eval_fps (f / g) z = eval_fps f z / eval_fps g z" proof - from nz[of 0] and \r > 0\ have nz': "fps_nth g 0 \ 0" by (auto simp: eval_fps_at_0 zero_ereal_def) have "R \ min r (fps_conv_radius g)" by (auto simp: R_def intro: min.coboundedI2) also have "min r (fps_conv_radius g) \ fps_conv_radius (inverse g)" by (intro fps_conv_radius_inverse assms) (auto simp: zero_ereal_def) finally have radius: "fps_conv_radius (inverse g) \ R" . have "R \ min (fps_conv_radius f) (fps_conv_radius (inverse g))" by (intro radius min.boundedI) (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) also have "\ \ fps_conv_radius (f * inverse g)" by (rule fps_conv_radius_mult) also have "f * inverse g = f / g" by (intro fps_divide_unit [symmetric] nz') finally show "fps_conv_radius (f / g) \ R" . assume z: "ereal (norm z) < R" have "eval_fps (f * inverse g) z = eval_fps f z * eval_fps (inverse g) z" using radius by (intro eval_fps_mult less_le_trans[OF z]) (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) also have "eval_fps (inverse g) z = inverse (eval_fps g z)" using \r > 0\ by (intro eval_fps_inverse[where r = r] less_le_trans[OF z] nz) (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) also have "f * inverse g = f / g" by fact finally show "eval_fps (f / g) z = eval_fps f z / eval_fps g z" by (simp add: field_split_simps) qed lemma fixes f g :: "complex fps" and r :: ereal defines "R \ Min {r, fps_conv_radius f, fps_conv_radius g}" assumes "subdegree g \ subdegree f" assumes "fps_conv_radius f > 0" "fps_conv_radius g > 0" "r > 0" assumes "\z. z \ eball 0 r \ z \ 0 \ eval_fps g z \ 0" shows fps_conv_radius_divide: "fps_conv_radius (f / g) \ R" and eval_fps_divide: "ereal (norm z) < R \ c = fps_nth f (subdegree g) / fps_nth g (subdegree g) \ eval_fps (f / g) z = (if z = 0 then c else eval_fps f z / eval_fps g z)" proof - define f' g' where "f' = fps_shift (subdegree g) f" and "g' = fps_shift (subdegree g) g" have f_eq: "f = f' * fps_X ^ subdegree g" and g_eq: "g = g' * fps_X ^ subdegree g" unfolding f'_def g'_def by (rule subdegree_decompose' le_refl | fact)+ have subdegree: "subdegree f' = subdegree f - subdegree g" "subdegree g' = 0" using assms(2) by (simp_all add: f'_def g'_def) have [simp]: "fps_conv_radius f' = fps_conv_radius f" "fps_conv_radius g' = fps_conv_radius g" by (simp_all add: f'_def g'_def) have [simp]: "fps_nth f' 0 = fps_nth f (subdegree g)" "fps_nth g' 0 = fps_nth g (subdegree g)" by (simp_all add: f'_def g'_def) have g_nz: "g \ 0" proof - define z :: complex where "z = (if r = \ then 1 else of_real (real_of_ereal r / 2))" from \r > 0\ have "z \ eball 0 r" by (cases r) (auto simp: z_def eball_def) moreover have "z \ 0" using \r > 0\ by (cases r) (auto simp: z_def) ultimately have "eval_fps g z \ 0" by (rule assms(6)) thus "g \ 0" by auto qed have fg: "f / g = f' * inverse g'" by (subst f_eq, subst (2) g_eq) (insert g_nz, simp add: fps_divide_unit) have g'_nz: "eval_fps g' z \ 0" if z: "norm z < min r (fps_conv_radius g)" for z proof (cases "z = 0") case False with assms and z have "eval_fps g z \ 0" by auto also from z have "eval_fps g z = eval_fps g' z * z ^ subdegree g" by (subst g_eq) (auto simp: eval_fps_mult) finally show ?thesis by auto qed (insert \g \ 0\, auto simp: g'_def eval_fps_at_0) have "R \ min (min r (fps_conv_radius g)) (fps_conv_radius g')" by (auto simp: R_def min.coboundedI1 min.coboundedI2) also have "\ \ fps_conv_radius (inverse g')" using g'_nz by (rule fps_conv_radius_inverse) finally have conv_radius_inv: "R \ fps_conv_radius (inverse g')" . hence "R \ fps_conv_radius (f' * inverse g')" by (intro order.trans[OF _ fps_conv_radius_mult]) (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) thus "fps_conv_radius (f / g) \ R" by (simp add: fg) fix z c :: complex assume z: "ereal (norm z) < R" assume c: "c = fps_nth f (subdegree g) / fps_nth g (subdegree g)" show "eval_fps (f / g) z = (if z = 0 then c else eval_fps f z / eval_fps g z)" proof (cases "z = 0") case False from z and conv_radius_inv have "ereal (norm z) < fps_conv_radius (inverse g')" by simp with z have "eval_fps (f / g) z = eval_fps f' z * eval_fps (inverse g') z" unfolding fg by (subst eval_fps_mult) (auto simp: R_def) also have "eval_fps (inverse g') z = inverse (eval_fps g' z)" using z by (intro eval_fps_inverse[of "min r (fps_conv_radius g')"] g'_nz) (auto simp: R_def) also have "eval_fps f' z * \ = eval_fps f z / eval_fps g z" using z False assms(2) by (simp add: f'_def g'_def eval_fps_shift R_def) finally show ?thesis using False by simp qed (simp_all add: eval_fps_at_0 fg field_simps c) qed lemma has_fps_expansion_fps_expansion [intro]: assumes "open A" "0 \ A" "f holomorphic_on A" shows "f has_fps_expansion fps_expansion f 0" proof - from assms(1,2) obtain r where r: "r > 0 " "ball 0 r \ A" by (auto simp: open_contains_ball) have holo: "f holomorphic_on eball 0 (ereal r)" using r(2) and assms(3) by auto from r(1) have "0 < ereal r" by simp also have "r \ fps_conv_radius (fps_expansion f 0)" using holo by (intro conv_radius_fps_expansion) auto finally have "\ > 0" . moreover have "eventually (\z. z \ ball 0 r) (nhds 0)" using r(1) by (intro eventually_nhds_in_open) auto hence "eventually (\z. eval_fps (fps_expansion f 0) z = f z) (nhds 0)" by eventually_elim (subst eval_fps_expansion'[OF holo], auto) ultimately show ?thesis using r(1) by (auto simp: has_fps_expansion_def) qed lemma fps_conv_radius_tan: fixes c :: complex assumes "c \ 0" shows "fps_conv_radius (fps_tan c) \ pi / (2 * norm c)" proof - have "fps_conv_radius (fps_tan c) \ Min {pi / (2 * norm c), fps_conv_radius (fps_sin c), fps_conv_radius (fps_cos c)}" unfolding fps_tan_def proof (rule fps_conv_radius_divide) fix z :: complex assume "z \ eball 0 (pi / (2 * norm c))" with cos_eq_zero_imp_norm_ge[of "c*z"] assms show "eval_fps (fps_cos c) z \ 0" by (auto simp: norm_mult field_simps) qed (insert assms, auto) thus ?thesis by (simp add: min_def) qed lemma eval_fps_tan: fixes c :: complex assumes "norm z < pi / (2 * norm c)" shows "eval_fps (fps_tan c) z = tan (c * z)" proof (cases "c = 0") case False show ?thesis unfolding fps_tan_def proof (subst eval_fps_divide'[where r = "pi / (2 * norm c)"]) fix z :: complex assume "z \ eball 0 (pi / (2 * norm c))" with cos_eq_zero_imp_norm_ge[of "c*z"] assms show "eval_fps (fps_cos c) z \ 0" using False by (auto simp: norm_mult field_simps) qed (use False assms in \auto simp: field_simps tan_def\) qed simp_all end