# HG changeset patch # User paulson # Date 1528753413 -3600 # Node ID 529d6b132c27b6cdf9c4b1cdbebb5772558bb4c0 # Parent 4e27f5c361d2674805b6a5f92cb891ea313048c9 tidier Cauchy proofs diff -r 4e27f5c361d2 -r 529d6b132c27 src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Sat Jun 09 21:52:16 2018 +0200 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon Jun 11 22:43:33 2018 +0100 @@ -6338,12 +6338,10 @@ have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" by (rule holomorphic_intros)+ show thesis - apply (rule Liouville_weak_inverse [OF 1]) - apply (rule polyfun_extremal) - apply (rule nz) - using i that - apply auto - done + 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 polyfun_extremal nz by force + qed (use that in auto) qed @@ -6358,14 +6356,15 @@ 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 continuous_on_sing intro: that) + 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 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" - apply (rule continuous_intros continuous_on_subset [OF contg])+ + 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 - @@ -6389,18 +6388,16 @@ 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" - apply (rule Lim_transform_eventually [where f = "\n. contour_integral (circlepath z r) (\u. f n u/(u - w))"]) - apply (rule eventually_mono [OF cont]) - apply (rule contour_integral_unique [OF Cauchy_integral_circlepath]) - using w - apply (auto simp: norm_minus_commute dist_norm cif_tends_cig) - done - have "((\n. 2 * of_real pi * \ * f n w) \ 2 * of_real pi * \ * g w) F" - apply (rule tendsto_mult_left [OF tendstoI]) - apply (rule eventually_mono [OF uniform_limitD [OF ulim]], assumption) - using w - apply (force simp: dist_norm) - done + 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 + 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 (force simp: dist_norm) @@ -6455,12 +6452,17 @@ 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: "\y. r = cmod (z - y) \ d \ cmod ((y - w)\<^sup>2)" - apply (simp add: d_def norm_power) - apply (rule power_mono) - apply (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) - apply (metis diff_ge_0_iff_ge dist_commute dist_norm less_eq_real_def mem_ball w) - done + 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" @@ -6468,9 +6470,8 @@ 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: diff_divide_distrib [symmetric] norm_divide divide_simps sphere_def dist_norm) + 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 divide_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) @@ -6508,10 +6509,12 @@ 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" - apply (intro eventuallyI conjI) - using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r apply blast - apply (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) - done + 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 apply (rule holomorphic_uniform_limit [OF *]) using \0 < r\ centre_in_ball ul @@ -6537,15 +6540,14 @@ 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))" - apply (intro eventuallyI conjI) - apply (meson hfd holomorphic_on_imp_continuous_on holomorphic_on_open holomorphic_on_subset r S) - using ball_subset_cball hfd r apply blast - done + 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 - apply (rule has_complex_derivative_uniform_limit [OF *, of g]) - using \0 < r\ centre_in_ball ul - apply force+ - done + 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) @@ -6569,11 +6571,11 @@ proof - obtain d where "d>0" and d: "cball x d \ S" using open_contains_cball [of "S"] \x \ S\ S by blast - then show ?thesis - apply (rule_tac x=d in exI) - using g uniform_limit_on_subset - apply (force simp: dist_norm eventually_sequentially) - done + 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]) @@ -6612,14 +6614,14 @@ 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)" - apply (rule_tac f=g in DERIV_transform_at [OF _ \0 < r\]) - apply (simp add: gg' \z \ S\ \0 < d\) - apply (metis (full_types) contra_subsetD dist_commute gg' mem_ball r sums_unique) - done + proof (rule DERIV_transform_at) + 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 + by (rule_tac x="\x. suminf (\n. f n x)" in exI) meson qed @@ -6682,12 +6684,9 @@ (\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 (simp add: dist_norm) 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_divide norm_power) - apply (rule_tac x=0 in exI) - apply (force simp: dist_norm intro!: mult_left_mono power_mono y_le) + 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) @@ -6742,12 +6741,10 @@ apply (auto simp: assms dist_norm) done qed - show ?thesis - apply (rule_tac x="g' w" in exI) - apply (rule DERIV_transform_at [where f=g and d="(r - norm(z - w))/2"]) - using w gg' [of w] - apply (auto simp: dist_norm) - done + have "(f has_field_derivative g' w) (at w)" + by (rule DERIV_transform_at [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 @@ -6755,10 +6752,8 @@ 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) - using holomorphic_power_series apply force - apply (rule power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) - apply force + apply (intro iffI ballI holomorphic_power_series, assumption+) + apply (force intro: power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) done corollary power_series_analytic: @@ -6791,102 +6786,104 @@ done lemma holomorphic_fun_eq_0_on_connected: - assumes holf: "f holomorphic_on s" and "open s" - and cons: "connected s" + 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" + and "z \ S" "w \ S" shows "f w = 0" proof - - have *: "\x e. \ \xa. (deriv ^^ xa) f x = 0; ball x e \ s\ - \ ball x e \ (\n. {w \ s. (deriv ^^ n) f w = 0})" - apply auto - apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv]) - apply (rule holomorphic_on_subset [OF holf], simp_all) - by (metis funpow_add o_apply) - have 1: "openin (subtopology euclidean s) (\n. {w \ s. (deriv ^^ n) f w = 0})" + 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) + with that show ?thesis by auto + qed + have 1: "openin (subtopology euclidean S) (\n. {w \ S. (deriv ^^ n) f w = 0})" apply (rule open_subset, force) - using \open s\ + using \open S\ apply (simp add: open_contains_ball Ball_def) apply (erule all_forward) using "*" by auto blast+ - have 2: "closedin (subtopology euclidean s) (\n. {w \ s. (deriv ^^ n) f w = 0})" + have 2: "closedin (subtopology euclidean 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\] . + 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" + have 3: "(\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 - using cons der \z \ s\ + 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 (drule_tac x="\n. {w \ S. (deriv ^^ n) f w = 0}" in spec) apply (auto simp: 1 2 3) done qed lemma holomorphic_fun_eq_on_connected: - assumes "f holomorphic_on s" "g holomorphic_on s" and "open s" "connected s" + 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" + and "z \ S" "w \ S" shows "f w = g w" - apply (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" s z, simplified]) - apply (intro assms holomorphic_intros) - using assms apply simp_all - apply (subst higher_deriv_diff, auto) - done +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" + 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" + and "z \ S" "w \ S" shows "f w = f z" - apply (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" s z, simplified]) - apply (intro assms holomorphic_intros) - using assms apply simp_all - apply (subst higher_deriv_diff) - apply (intro holomorphic_intros | simp)+ - done - +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\Some basic lemmas about poles/singularities\ lemma pole_lemma: - assumes holf: "f holomorphic_on s" and a: "a \ interior s" + 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") + 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 + 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" + 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\) + 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 + 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\]) + by (rule holomorphic_on_subset [OF holf \ball a e \ S\]) have 2: "?F holomorphic_on ball a e - {a}" - apply (rule holomorphic_on_subset [where s = "s - {a}"]) - apply (simp add: holomorphic_on_def field_differentiable_def [symmetric]) + apply (simp add: holomorphic_on_def flip: field_differentiable_def) using mem_ball that apply (auto intro: F1 field_differentiable_within_subset) done 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 show ?thesis - using holfb \0 < e\ - apply (simp add: holomorphic_on_open field_differentiable_def [symmetric]) - apply (drule_tac x=a in bspec) - apply (auto simp: DERIV_deriv_iff_field_differentiable [symmetric] continuous_at DERIV_iff2 + 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 DERIV_iff2 simp flip: DERIV_deriv_iff_field_differentiable elim: rev_iffD1 [OF _ LIM_equal]) - done 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) @@ -6901,29 +6898,29 @@ qed show ?thesis proof - fix x assume "x \ s" show "?F field_differentiable at x within s" + 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\ + case False with F1 \x \ S\ show ?thesis by blast qed qed qed proposition 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" + 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" + else f z - g a/(z - a)) holomorphic_on S" using pole_lemma [OF holg a] by (rule holomorphic_transform) (simp add: eq divide_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") + 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 @@ -6935,48 +6932,53 @@ qed proposition 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" + 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] + 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) proposition 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" + 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" + shows "f holomorphic_on S" using pole_theorem [OF holg a eq] by (rule holomorphic_transform) (auto simp: eq divide_simps) proposition 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" + 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] + shows "f holomorphic_on S" + using pole_theorem_open [OF holg S eq] by (rule holomorphic_transform) (auto simp: eq divide_simps) lemma pole_theorem_analytic: - assumes g: "g analytic_on s" - and eq: "\z. z \ s + 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" -using g -apply (simp add: analytic_on_def Ball_def) -apply (safe elim!: all_forward dest!: eq) -apply (rule_tac x="min d e" in exI, simp) -apply (rule pole_theorem_open) -apply (auto simp: holomorphic_on_subset subset_ball) -done + 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)" + 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" + 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 @@ -6985,22 +6987,27 @@ 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" + 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" - apply (rule pole_theorem_analytic [OF g]) - apply (rule openE [OF \open t\]) - using assms eq by auto + 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" + 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" + 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" + 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 @@ -7011,24 +7018,25 @@ 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))" + 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 \ + 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 \ + 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 = "{(t,z) |t z. t \ cball w \ \ z \ closed_segment a b}" + 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)" - apply (rule compact_uniformly_continuous) - apply (rule continuous_on_subset[OF cond_uu]) - using abu \ - apply (auto simp: compact_Times simp del: mem_cball) - done + 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)" @@ -7040,13 +7048,13 @@ 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' + if "x' \ U" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' proof - - have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" + 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 _ \]) - apply (rule contour_integrable_diff [OF cont_dw cont_dw]) - using \0 < \\ \a \ b\ \0 < \\ \w \ u\ that - apply (auto simp: norm_minus_commute) + using \0 < \\ \0 < \\ that apply (auto simp: norm_minus_commute) done also have "\ = \" using \a \ b\ by simp finally show ?thesis . @@ -7054,22 +7062,26 @@ 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) + apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \w \ U\ intro: le_ee) done qed show ?thesis - apply (rule continuous_onI) - apply (cases "a=b") - apply (auto intro: *) - done + 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 text\This version has @{term"polynomial_function \"} as an additional assumption.\ lemma Cauchy_integral_formula_global_weak: - assumes u: "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" + 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)" @@ -7084,46 +7096,50 @@ 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" + have uv_Un: "U \ v = UNIV" using pasz zero by (auto simp: v_def) - have conf: "continuous_on u f" + 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 + 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 u) + 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 u by fastforce - then have "continuous_on u (d y)" - apply (simp add: d_def continuous_on_eq_continuous_at u, clarify) + 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 u) - moreover have "d y holomorphic_on u - {y}" - apply (simp add: d_def holomorphic_on_open u open_delete field_differentiable_def [symmetric]) - apply (intro ballI) - apply (rename_tac 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 analytic_on_imp_differentiable_at analytic_on_open holf u apply blast - done + by (meson field_differentiable_within_open field_differentiable_imp_continuous_at \open U\) + 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 + 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: u) + 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 - apply (rule contour_integrable_holomorphic_simple [where S = "u-{y}"]) - using \valid_path \\ pasz - apply (auto simp: u open_delete) - apply (rule continuous_intros holomorphic_intros continuous_on_subset [OF conf] holomorphic_on_subset [OF holf] | - force simp: that)+ - done + 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: "\z. z \ u \ ((d z) has_contour_integral h z) \" + "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 apply (simp add: h_def) - apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where S=u]]) - using u pasz \valid_path \\ - apply (auto intro: holomorphic_on_imp_continuous_on hol_d) - done + by (meson Diff_subset \open U\ \valid_path \\ contour_integrable_holomorphic_simple 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" @@ -7142,24 +7158,24 @@ 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 + 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 | force)+ + apply (rule continuous_intros conf holomorphic_intros holf assms | force)+ done 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 u + 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\ by (fastforce simp add: \path \\ compact_path_image) - obtain dd where "0 < dd" and dd: "{y + k | y k. y \ path_image \ \ k \ ball 0 dd} \ u" + 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) @@ -7174,27 +7190,27 @@ using \0 < dd\ 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" + 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 obtain L where "L>0" - and L: "\f B. \f holomorphic_on interior t; \z. z\interior t \ cmod (f z) \ B\ \ + 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" + 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 + 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 + 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" @@ -7204,12 +7220,12 @@ done 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" + 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 + using holf holomorphic_on_subset interior_subset T apply blast apply (rule holomorphic_intros)+ - using \y \ t\ interior_subset by auto - have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior t" for z + using \y \ T\ interior_subset by 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 @@ -7238,32 +7254,33 @@ 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 + if "x \ U" "z \ U" "x \ z" for x z using that conf - apply (simp add: split_def continuous_on_eq_continuous_at u) + 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)" + have open_uu_Id: "open (U \ U - Id)" apply (rule open_Diff) - apply (simp add: open_Times u) + 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 - 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 u) - by (simp add: interior_open that u) + 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\) 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 + (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\]] + 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 u] \x \ u\) + 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' @@ -7273,9 +7290,9 @@ 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 u) - have "closed_segment x' z' \ u" + 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 @@ -7290,7 +7307,7 @@ also have "\ \ e" using \0 < e\ by simp finally show ?thesis . qed - show "\d>0. \xa\u \ u. + 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) @@ -7299,49 +7316,51 @@ done 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" + 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 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)" + 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 u, clarify) + 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 + 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)))" + 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))" + 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" + 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 u derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp: dist_commute)+ + apply (rule \open U\ derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp: dist_commute)+ done show ?thesis unfolding d_def - apply (rule no_isolated_singularity [OF * _ u, where K = "{w}"]) - apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff u **) + 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 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)" + 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]) + 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 {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) \ \)" - apply (rule continuous_on_compose) - using \path \\ path_def pasz - apply (auto intro!: continuous_on_subset [OF cont_cint_d]) - apply (force simp: path_image_def) - done + 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 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) @@ -7361,13 +7380,13 @@ 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" + 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" + 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 u x by force + 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 @@ -7382,10 +7401,9 @@ apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) done then obtain kk where "kk>0" - and kk: "\x x'. \x\?ddpa; x'\?ddpa; dist x' x < kk\ \ + 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" - apply (rule uniformly_continuous_onE [where e = ee]) - using \0 < ee\ by auto + 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) @@ -7397,35 +7415,34 @@ done qed qed - have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" - apply (simp add: contour_integral_unique [OF U, symmetric] x) - apply (rule contour_integral_uniform_limit [OF A1 A2 le_B]) - apply (auto simp: \valid_path \\) - done + 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 + 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 u]) + 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 u + 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" + 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" + 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 @@ -7434,7 +7451,7 @@ (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 + using abc pasz U apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ done then show ?thesis @@ -7540,13 +7557,12 @@ "path_image g \ S" "pathfinish g = pathstart g" "z \ S" shows "winding_number g z = 0" proof - - have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z" - apply (rule winding_number_homotopic_paths) - apply (rule homotopic_loops_imp_homotopic_paths_null [where a = "pathstart g"]) - apply (rule homotopic_loops_subset [of S]) - using assms - apply (auto simp: homotopic_paths_imp_homotopic_loops path_defs simply_connected_eq_contractible_path) - done + 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 . @@ -7562,7 +7578,7 @@ homotopic_paths_imp_homotopic_loops) using valid_path_imp_path by blast -lemma holomorphic_logarithm_exists: +proposition 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" @@ -7586,7 +7602,6 @@ 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"