# HG changeset patch # User paulson # Date 1600364917 -3600 # Node ID 121b838a0ba891a93d676d1877ba68bf0b61a57f # Parent ff32ddc8165c7d584e5f5cb5a72ba7053813bfb8# Parent 1e02b86eb517ba472dba1c73f31230eb526f860d merged diff -r ff32ddc8165c -r 121b838a0ba8 src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy --- a/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy Thu Sep 17 14:27:56 2020 +0200 +++ b/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy Thu Sep 17 18:48:37 2020 +0100 @@ -6,45 +6,43 @@ 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 \" + 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 + 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 - apply (simp add: continuous_within) - apply (rule Lim_transform_away_within [of _ "z+1" _ "\w::complex. (f w - f z)/(w - z)"]) - using has_field_derivative_at_within has_field_derivative_iff f' - apply (fastforce simp add:)+ - done + 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" + 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))" + 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 "\w::complex. (f w - f z)/(w - z)"]) + 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 (f w - f z) / (w - z)) has_contour_integral 0) \" - apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop]) - using c apply (force simp: continuous_on_eq_continuous_within) - apply (rename_tac w) - apply (rule_tac d="dist w z" and f = "\w. (f w - f z)/(w - z)" in field_differentiable_transform_within) - apply (simp_all add: dist_pos_lt dist_commute) - apply (metis less_irrefl) - apply (rule derivative_intros fcd | simp)+ - done + 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"] *] @@ -53,12 +51,16 @@ qed theorem Cauchy_integral_formula_convex_simple: - "\convex s; f holomorphic_on s; z \ interior s; valid_path \; path_image \ \ s - {z}; - pathfinish \ = pathstart \\ - \ ((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" - apply (rule Cauchy_integral_formula_weak [where k = "{}"]) - using holomorphic_on_imp_continuous_on - by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE) + 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.\ @@ -71,7 +73,7 @@ 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 = "{}"]) + 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 @@ -95,25 +97,25 @@ 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 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 "open S" and \: "valid_path \" - and w: "w \ s - path_image \" + 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 + 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 \" - apply (rule contour_integrable_div [OF contour_integrable_diff]) using int w d - by (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+ + 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 @@ -260,13 +262,14 @@ (f u - f w) / (u - w) / k" if "dist u w < d" for u proof - - have u: "u \ s - path_image \" + 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) - apply (intro has_contour_integral_diff has_contour_integral_div) - using u w apply (simp_all add: field_simps int) + apply (simp add: diff_divide_distrib algebra_simps \
has_contour_integral_diff has_contour_integral_div) done qed show ?thes2 @@ -342,8 +345,8 @@ 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" - apply (simp add: holomorphic_on_def) - using field_differentiable_at_within field_differentiable_def fder by blast + 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)" @@ -372,10 +375,8 @@ by (simp add: f'_eq) } note * = this show ?thesis - apply (rule exI) - apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) - apply (simp_all add: \0 < r\ * dist_norm) - done + 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\] *) @@ -434,18 +435,19 @@ contour_integral (linepath c a) f = 0" by blast have az: "dist a z < e" using mem_ball z by blast - have sb_ball: "ball z (e - dist a z) \ ball a e" - by (simp add: dist_commute ball_subset_ball_iff) have "\e>0. f holomorphic_on ball z e" proof (intro exI conjI) - 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 z (e - dist a z)" - apply (rule holomorphic_on_subset [OF _ sb_ball]) - apply (rule derivative_is_holomorphic[OF open_ball]) - apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]) - apply (simp_all add: 0 \0 < e\ sub_ball) - done + 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 @@ -507,9 +509,10 @@ lemma higher_deriv_ident [simp]: "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" - apply (induction n, simp) - apply (metis higher_deriv_linear lambda_one) - done +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)" @@ -517,9 +520,15 @@ lemma has_complex_derivative_funpow_1: "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" - apply (induction n, auto) - apply (simp add: id_def) - by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) +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" @@ -531,11 +540,11 @@ 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 "((deriv ^^ n) (\w. - f w) has_field_derivative - deriv ((deriv ^^ n) f) z) (at z)" - apply (rule has_field_derivative_transform_within_open [of "\w. -((deriv ^^ n) f w)"]) - apply (rule derivative_eq_intros | rule * refl assms)+ - apply (auto simp add: Suc) - done + 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 @@ -552,24 +561,22 @@ 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 "((deriv ^^ n) (\w. f w + g w) has_field_derivative + 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)" - apply (rule has_field_derivative_transform_within_open [of "\w. (deriv ^^ n) f w + (deriv ^^ n) g w"]) - apply (rule derivative_eq_intros | rule * refl assms)+ - apply (auto simp add: Suc) - done + 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" and z: "z \ S" + 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" - apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add) - apply (subst higher_deriv_add) - using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus) - done + 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 @@ -598,12 +605,11 @@ (\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)"]) + [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 DERIV_cong [OF DERIV_sum]) - apply (rule DERIV_cmult) - apply (auto intro: DERIV_mult * sumeq \open S\ Suc.prems Suc.IH [symmetric]) - done + 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) @@ -634,29 +640,36 @@ 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 holo1: "(\w. f (u * w)) holomorphic_on S" - apply (rule holomorphic_on_compose [where g=f, unfolded o_def]) - apply (rule holo0 holomorphic_intros)+ - done + 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" - apply (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems]) - apply (rule holomorphic_higher_deriv [OF holo1 S]) - apply (simp add: Suc.IH) - done + 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" - apply (rule deriv_cmult) - apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) - apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and T=T, unfolded o_def]) - apply (simp) - apply (simp add: analytic_on_open f holomorphic_higher_deriv T) - apply (blast intro: fg) - done + 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)" - apply (subst deriv_chain [where g = "(deriv ^^ n) f" and f = "(*) u", unfolded o_def]) - apply (rule derivative_intros) - using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv T apply blast - apply (simp) - done + 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 @@ -896,8 +909,8 @@ if "n \ N" and r: "r = dist z u" for N u proof - have N: "((r - k) / r) ^ N < e / B * k" - apply (rule le_less_trans [OF power_decreasing n]) - using \n \ N\ k by auto + 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" @@ -918,10 +931,7 @@ by (simp add: norm_mult norm_power norm_minus_commute) also have "\ \ (((r - k)/r)^N) * B" using \0 < r\ w k - apply (simp add: divide_simps) - apply (rule mult_mono [OF power_mono]) - apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r) - done + 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)" @@ -933,23 +943,24 @@ 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) - using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps) - apply (simp only: contour_integral_lmul cint algebra_simps) + apply (simp_all only: \
contour_integral_lmul cint algebra_simps) done - have cic: "\u. (\y. \ku 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]) - have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) + 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 - apply (intro Lim_transform_eventually [OF _ eq] contour_integral_uniform_limit_circlepath [OF eventuallyI ul] cic) - using \0 < r\ apply auto - done + 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]]) @@ -984,9 +995,8 @@ by linarith qed have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" - apply (rule Cauchy_integral_circlepath) - using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ - done + 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) @@ -1009,10 +1019,12 @@ 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" - apply (rule tendstoI [OF eventually_mono]) - apply (rule_tac B="2/e" in unbounded) - apply (simp add: dist_norm norm_divide field_split_simps) - done + 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 } @@ -1070,16 +1082,19 @@ 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" - apply (rule eventually_mono [OF cont]) using w - apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified]) - done - 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\ - apply (clarsimp simp add: uniform_limit_iff dist_norm norm_divide diff_divide_distrib [symmetric] divide_simps) + 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" @@ -1088,8 +1103,8 @@ 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" - apply (rule eventually_mono [OF cont contour_integral_unique [OF Cauchy_integral_circlepath]]) - using w\0 < d\ d_def by auto + 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) @@ -1166,13 +1181,13 @@ unfolding uniform_limit_iff proof clarify fix e::real - assume "0 < e" - with \r > 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" - apply (simp add: norm_divide field_split_simps sphere_def dist_norm) - apply (rule eventually_mono [OF uniform_limitD [OF ulim], of "e*d"]) - apply (simp add: \0 < d\) - apply (force simp: dist_norm dle intro: less_le_trans) - done + 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" @@ -1213,10 +1228,8 @@ by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) qed show ?thesis - apply (rule holomorphic_uniform_limit [OF *]) using \0 < r\ centre_in_ball ul - apply (auto simp: holomorphic_on_open) - done + by (auto simp: holomorphic_on_open intro: holomorphic_uniform_limit [OF *]) qed with S show ?thesis by (simp add: holomorphic_on_open) @@ -1363,12 +1376,16 @@ 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 (z - y) * 2 < r - cmod z\ \ cmod y \ cmod (of_real r + of_real (cmod z)) / 2" for z y - using \r > 0\ - apply (auto simp: algebra_simps norm_mult norm_divide norm_power simp flip: of_real_add) - using norm_triangle_ineq2 [of y z] - apply (simp only: diff_le_eq norm_minus_commute mult_2) - done + 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)" @@ -1388,7 +1405,7 @@ by (simp add: ball_def) next case False then show ?thesis - apply (simp add: not_less) + unfolding not_less using less_le_trans norm_not_less_zero by blast qed @@ -1431,11 +1448,10 @@ 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) - show ?thesis - apply (rule sums_unique2 [of "\n. a n*(u - z)^n"]) - using gg' [of u] less w - apply (auto simp: assms dist_norm) - done + 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"]) @@ -1468,17 +1484,13 @@ w \ ball z r; \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ \ f w = g w" - apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) - apply (auto simp: holomorphic_iff_power_series) - done + 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" - apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) - apply (auto simp: holomorphic_iff_power_series) - done + 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" @@ -1490,33 +1502,31 @@ 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 "\x' n. dist x x' < e \ (deriv ^^ n) f x' = 0" - apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv]) - apply (rule holomorphic_on_subset [OF holf]) - using that apply simp_all - by (metis funpow_add o_apply) + 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 - have 1: "openin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" - apply (rule open_subset, force) + 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+ - have 2: "closedin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" + 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) - 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 3: "(\n. {w \ S. (deriv ^^ n) f w = 0}) = S \ f w = 0" + 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]) - show ?thesis + ultimately show ?thesis using cons der \z \ S\ - apply (simp add: connected_clopen) - apply (drule_tac x="\n. {w \ S. (deriv ^^ n) f w = 0}" in spec) - apply (auto simp: 1 2 3) - done + by (auto simp add: connected_clopen) qed lemma holomorphic_fun_eq_on_connected: @@ -1566,10 +1576,8 @@ 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}" - apply (simp add: holomorphic_on_def flip: field_differentiable_def) using mem_ball that - apply (auto intro: F1 field_differentiable_within_subset) - done + 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") @@ -1735,8 +1743,8 @@ 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)" - apply (rule uniformly_continuous_onE [where e = "\/norm(b - a)"]) - using \0 < \\ \a \ b\ by auto + 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)" @@ -1749,26 +1757,21 @@ 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 apply (auto simp: norm_minus_commute) - done + 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 < \\ - apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \w \ U\ intro: le_ee) - done + 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 True - then show ?thesis by simp - next case False show ?thesis by (rule continuous_onI) (use False in \auto intro: *\) - qed + qed auto qed text\This version has \<^term>\polynomial_function \\ as an additional assumption.\ @@ -1802,16 +1805,18 @@ 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)" - apply (simp add: d_def continuous_on_eq_continuous_at \open U\, clarify) - using * holomorphic_on_def - by (meson field_differentiable_within_open field_differentiable_imp_continuous_at \open U\) + using "*" d_def holomorphic_on_imp_continuous_on by auto moreover have "d y holomorphic_on U - {y}" proof - - have "\w. w \ U - {y} \ - (\w. if w = y then deriv f y else (f w - f y) / (w - y)) field_differentiable at w" - apply (rule_tac d="dist w y" and f = "\w. (f w - f y)/(w - y)" in field_differentiable_transform_within) - apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros) - using \open U\ holf holomorphic_on_imp_differentiable_at by blast + 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 @@ -1832,8 +1837,7 @@ have "d z holomorphic_on U" by (simp add: hol_d that) with that show ?thesis - apply (simp add: h_def) - by (meson Diff_subset \open U\ \valid_path \\ contour_integrable_holomorphic_simple has_contour_integral_integral pasz subset_trans) + 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 - @@ -1847,21 +1851,20 @@ 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 (auto simp: v_def) + 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" + 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" - apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where S=U]]) - using U pasz \valid_path \\ that - apply (auto intro: holomorphic_on_imp_continuous_on hol_d) - apply (rule continuous_intros conf holomorphic_intros holf assms | force)+ - done + 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 @@ -1871,25 +1874,25 @@ 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" - apply (rule that [of "d0/2"]) - using \0 < d0\ - apply (auto simp: dist_norm dest: d0) - done + 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 1: "path_image \ \ interior {y + k |y k. y \ path_image \ \ k \ cball 0 (dd / 2)}" - apply (clarsimp simp add: mem_interior) - using \0 < dd\ + 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 - obtain T where "compact T" and subt: "path_image \ \ interior T" and T: "T \ U" - apply (rule that [OF _ 1]) - apply (fastforce simp add: \valid_path \\ compact_valid_path_image intro!: compact_sums) - apply (rule order_trans [OF _ dd]) - using \0 < dd\ by fastforce + 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" @@ -1909,17 +1912,17 @@ then have ynot: "y \ path_image \" using subt interior_subset by blast have [simp]: "winding_number \ y = 0" - apply (rule winding_number_zero_outside [of _ "cball 0 C"]) - using ybig interior_subset subt - apply (force simp: loop \path \\ dist_norm intro!: C)+ - done + 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" - apply (rule holomorphic_on_divide) - using holf holomorphic_on_subset interior_subset T apply blast - apply (rule holomorphic_intros)+ - using \y \ T\ interior_subset by auto + 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" @@ -1929,11 +1932,12 @@ 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)" - apply (rule mult_mono) - using that D interior_subset apply blast - using \L>0\ \e>0\ \D>0\ DL2 - apply (auto simp: norm_divide field_split_simps) - done + 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)))" @@ -1957,15 +1961,18 @@ have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" by (rule continuous_intros)+ have open_uu_Id: "open (U \ U - Id)" - apply (rule open_Diff) - apply (simp add: open_Times \open U\) - using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] - apply (auto simp: Id_fstsnd_eq algebra_simps) - done + 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 - apply (rule continuous_on_interior [of U]) - apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open U\) - by (simp add: interior_open that \open U\) + 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 @@ -1981,7 +1988,7 @@ for x' z' proof - have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w - apply (drule segment_furthest_le [where y=x]) + 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) @@ -2007,14 +2014,12 @@ 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\ - apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) - done + 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" - apply (rule B) - using \' using path_image_def vector_derivative_at by fastforce + 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)" @@ -2030,24 +2035,29 @@ 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 **: "\x. \x \ U; x \ w\ \ (\z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" - apply (rule_tac f = "\x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) - apply (rule \open U\ derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp: dist_commute)+ - done + 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 - apply (rule no_isolated_singularity [OF * _ \open U\, where K = "{w}"]) - apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff \open U\ **) - done + 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" - then have "\w. w \ U \ (\z. d z w) contour_integrable_on (linepath a b)" - by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) - then have cont_cint_d: "continuous_on U (\w. contour_integral (linepath a b) (\z. d z w))" - apply (rule contour_integral_continuous_on_linepath_2D [OF \open U\ _ _ abu]) - apply (auto intro: continuous_on_swap_args cond_uu) - done + 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} \" @@ -2056,20 +2066,23 @@ using pasz unfolding path_image_def by (auto intro!: continuous_on_subset [OF cont_cint_d]) qed - have cint_cint: "(\w. contour_integral (linepath a b) (\z. d z w)) contour_integrable_on \" + 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) - apply (rule continuous_on_mult [OF cont_cint_d\ [unfolded o_def]]) - using pf\' - by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \']) + 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))" - apply (rule contour_integral_swap) - apply (rule continuous_on_subset [OF cond_uu]) - using abu pasz \valid_path \\ - apply (auto intro!: continuous_intros) - by (metis \' continuous_on_eq path_def path_polynomial_function pf\' vector_derivative_at) + 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))" . @@ -2091,10 +2104,11 @@ proof - let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" - apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) - using dd pasz \valid_path \\ - apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) - done + 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" @@ -2102,12 +2116,11 @@ 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) - show ?thesis - using ax unfolding lim_sequentially eventually_sequentially - apply (drule_tac x="min dd kk" in spec) - using \dd > 0\ \kk > 0\ - apply (fastforce simp: kk dist_norm) - done + 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)" @@ -2138,9 +2151,10 @@ 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" - apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) - apply (rule holomorphic_on_subset [OF hol_dw]) - using e abc_subset by auto + 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) + @@ -2316,12 +2330,13 @@ proof - have "0 < B0" using \0 < r\ fin [of z] by (metis ball_eq_empty ex_in_conv fin not_less) - have le_B0: "\w. cmod (w - z) \ r \ cmod (f w - y) \ B0" - apply (rule continuous_on_closure_norm_le [of "ball z r" "\w. f w - y"]) - apply (auto simp: \0 < r\ dist_norm norm_minus_commute) - apply (rule continuous_intros contf)+ - using fin apply (simp add: dist_commute dist_norm less_eq_real_def) - done + have 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" @@ -2362,15 +2377,22 @@ dual_order.strict_implies_order norm_of_real) then have "0 \ B" by (metis nof norm_not_less_zero not_le order_trans) - have "((\u. f u / (u - \) ^ Suc n) has_contour_integral (2 * pi) * \ / fact n * (deriv ^^ n) f \) - (circlepath \ r)" - apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) + have "\ \ ball \ r" using \0 < r\ by simp - then have "norm ((2 * pi * \)/(fact n) * (deriv ^^ n) f \) \ (B / r^(Suc n)) * (2 * pi * r)" - apply (rule has_contour_integral_bound_circlepath) - using \0 \ B\ \0 < r\ - apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc) - done + then 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 @@ -2392,9 +2414,10 @@ next assume "0 < B" have "((\k. (deriv ^^ k) f 0 / (fact k) * (\ - 0)^k) sums f \)" - apply (rule holomorphic_power_series [where r = "norm \ + 1"]) - using holf holomorphic_on_subset apply auto - done + 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") @@ -2420,15 +2443,21 @@ 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" - apply (rule Cauchy_inequality) - using holf holomorphic_on_subset apply force - using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast - using \w \ 0\ apply simp - by (metis nof wgeA dist_0_norm dist_norm) + 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))" - apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) - using \k>n\ \w \ 0\ \0 < B\ apply (simp add: field_split_simps semiring_normalization_rules) - done + 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)" . @@ -2466,7 +2495,7 @@ 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 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" @@ -2474,24 +2503,23 @@ 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)" - apply (rule_tac m = "LEAST n. a n \ 0" in that) - using m0 - apply (rule LeastI2) - apply (fastforce intro: dest!: not_less_Least)+ - done + 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 + 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"] + 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))" + 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 + 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 \=\]) @@ -2499,18 +2527,17 @@ using not_le order_trans r by fastforce then have "continuous_on (cball \ (r/2)) g" using powser_continuous_suminf [of "r/2" b \] by (simp add: g_def) - then obtain s where "s>0" "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ dist (g x) (g \) < 1/2" - apply (rule continuous_onE [where x=\ and e = "1/2"]) - using r apply (auto simp: norm_minus_commute dist_norm) - done + 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" + ultimately have gnz: "\x. \norm (x-\) \ s; norm (x-\) \ r/2\ \ (g x) \ 0" by fastforce - have "f x \ 0" if "x \ \" "norm (x - \) \ s" "norm (x - \) \ r/2" for x - using bsums [of x] that gnz [of x] - apply (auto simp: g_def) - using r sums_iff by fastforce + 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) @@ -2780,8 +2807,8 @@ 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 (insert False assms, auto simp: field_simps tan_def) - qed simp_all + show "eval_fps (fps_cos c) z \ 0" using False by (auto simp: norm_mult field_simps) + qed (insert False assms, auto simp: field_simps tan_def) +qed simp_all end