diff --git a/src/HOL/Analysis/Cauchy_Integral_Formula.thy b/src/HOL/Analysis/Cauchy_Integral_Formula.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Analysis/Cauchy_Integral_Formula.thy @@ -0,0 +1,2090 @@ +section \Cauchy's Integral Formula\ + +theory Cauchy_Integral_Formula + imports Winding_Numbers +begin + +subsection\Cauchy's integral formula, again for a convex enclosing set\ + +lemma Cauchy_integral_formula_weak: + assumes s: "convex s" and "finite k" and conf: "continuous_on s f" + and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" + and z: "z \ interior s - k" and vpg: "valid_path \" + and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + obtain f' where f': "(f has_field_derivative f') (at z)" + using fcd [OF z] by (auto simp: field_differentiable_def) + have pas: "path_image \ \ s" and znotin: "z \ path_image \" using pasz by blast+ + have c: "continuous (at x within s) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ s" for x + proof (cases "x = z") + case True then show ?thesis + 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 + next + case False + then have dxz: "dist x z > 0" by auto + have cf: "continuous (at x within s) f" + using conf continuous_on_eq_continuous_within that by blast + have "continuous (at x within s) (\w. (f w - f z) / (w - z))" + by (rule cf continuous_intros | simp add: False)+ + then show ?thesis + apply (rule continuous_transform_within [OF _ dxz that, of "\w::complex. (f w - f z)/(w - z)"]) + 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 + show ?thesis + apply (rule has_contour_integral_eq) + using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] + apply (auto simp: ac_simps divide_simps) + done +qed + +text\ Hence the Cauchy formula for points inside a circle.\ + +theorem Cauchy_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r" + shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) + (circlepath z r)" +proof - + have "r > 0" + using assms le_less_trans norm_ge_zero by blast + have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) + (circlepath z r)" + proof (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) + show "\x. x \ interior (cball z r) - {} \ + f field_differentiable at x" + using holf holomorphic_on_imp_differentiable_at by auto + have "w \ sphere z r" + by simp (metis dist_commute dist_norm not_le order_refl wz) + then show "path_image (circlepath z r) \ cball z r - {w}" + using \r > 0\ by (auto simp add: cball_def sphere_def) + qed (use wz in \simp_all add: dist_norm norm_minus_commute contf\) + then show ?thesis + by (simp add: winding_number_circlepath assms) +qed + +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) + +corollary\<^marker>\tag unimportant\ Cauchy_integral_circlepath_simple: + assumes "f holomorphic_on cball z r" "norm(w - z) < r" + shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) + (circlepath z r)" +using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) + +text\ In particular, the first derivative formula.\ + +lemma Cauchy_derivative_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" + (is "?thes2") +proof - + have [simp]: "r \ 0" using w + using ball_eq_empty by fastforce + have f: "continuous_on (path_image (circlepath z r)) f" + by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def) + have int: "\w. dist z w < r \ + ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" + by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) + show ?thes1 + apply (simp add: power2_eq_square) + apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) + apply (blast intro: int) + done + have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" + apply (simp add: power2_eq_square) + apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) + apply (blast intro: int) + done + then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" + by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) + show ?thes2 + by simp (rule fder) +qed + + +proposition derivative_is_holomorphic: + assumes "open S" + and fder: "\z. z \ S \ (f has_field_derivative f' z) (at z)" + shows "f' holomorphic_on S" +proof - + have *: "\h. (f' has_field_derivative h) (at z)" if "z \ S" for z + proof - + obtain r where "r > 0" and r: "cball z r \ S" + using open_contains_cball \z \ S\ \open S\ by blast + then have holf_cball: "f holomorphic_on cball z r" + apply (simp add: holomorphic_on_def) + using field_differentiable_at_within field_differentiable_def fder by blast + then have "continuous_on (path_image (circlepath z r)) f" + using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) + then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" + by (auto intro: continuous_intros)+ + have contf_cball: "continuous_on (cball z r) f" using holf_cball + by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) + have holf_ball: "f holomorphic_on ball z r" using holf_cball + using ball_subset_cball holomorphic_on_subset by blast + { fix w assume w: "w \ ball z r" + have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" + by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) + have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) + (at w)" + by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) + have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" + using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) + have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral + contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) + (circlepath z r)" + by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) + then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral + contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) + (circlepath z r)" + by (simp add: algebra_simps) + then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" + by (simp add: f'_eq) + } note * = this + show ?thesis + apply (rule exI) + apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) + apply (simp_all add: \0 < r\ * dist_norm) + done + qed + show ?thesis + by (simp add: holomorphic_on_open [OF \open S\] *) +qed + + +subsection\Existence of all higher derivatives\ + +lemma holomorphic_deriv [holomorphic_intros]: + "\f holomorphic_on S; open S\ \ (deriv f) holomorphic_on S" + by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) + +lemma analytic_deriv [analytic_intros]: "f analytic_on S \ (deriv f) analytic_on S" + using analytic_on_holomorphic holomorphic_deriv by auto + +lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv ^^ n) f holomorphic_on S" + by (induction n) (auto simp: holomorphic_deriv) + +lemma analytic_higher_deriv [analytic_intros]: "f analytic_on S \ (deriv ^^ n) f analytic_on S" + unfolding analytic_on_def using holomorphic_higher_deriv by blast + +lemma has_field_derivative_higher_deriv: + "\f holomorphic_on S; open S; x \ S\ + \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" +by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply + funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) + +lemma valid_path_compose_holomorphic: + assumes "valid_path g" and holo:"f holomorphic_on S" and "open S" "path_image g \ S" + shows "valid_path (f \ g)" +proof (rule valid_path_compose[OF \valid_path g\]) + fix x assume "x \ path_image g" + then show "f field_differentiable at x" + using analytic_on_imp_differentiable_at analytic_on_open assms holo by blast +next + have "deriv f holomorphic_on S" + using holomorphic_deriv holo \open S\ by auto + then show "continuous_on (path_image g) (deriv f)" + using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto +qed + +proposition\<^marker>\tag unimportant\ holomorphic_logarithm_exists: + assumes A: "convex A" "open A" + and f: "f holomorphic_on A" "\x. x \ A \ f x \ 0" + and z0: "z0 \ A" + obtains g where "g holomorphic_on A" and "\x. x \ A \ exp (g x) = f x" +proof - + note f' = holomorphic_derivI [OF f(1) A(2)] + obtain g where g: "\x. x \ A \ (g has_field_derivative deriv f x / f x) (at x)" + proof (rule holomorphic_convex_primitive' [OF A]) + show "(\x. deriv f x / f x) holomorphic_on A" + by (intro holomorphic_intros f A) + qed (auto simp: A at_within_open[of _ A]) + define h where "h = (\x. -g z0 + ln (f z0) + g x)" + from g and A have g_holo: "g holomorphic_on A" + by (auto simp: holomorphic_on_def at_within_open[of _ A] field_differentiable_def) + hence h_holo: "h holomorphic_on A" + by (auto simp: h_def intro!: holomorphic_intros) + have "\c. \x\A. f x / exp (h x) - 1 = c" + proof (rule has_field_derivative_zero_constant, goal_cases) + case (2 x) + note [simp] = at_within_open[OF _ \open A\] + from 2 and z0 and f show ?case + by (auto simp: h_def exp_diff field_simps intro!: derivative_eq_intros g f') + qed fact+ + then obtain c where c: "\x. x \ A \ f x / exp (h x) - 1 = c" + by blast + from c[OF z0] and z0 and f have "c = 0" + by (simp add: h_def) + with c have "\x. x \ A \ exp (h x) = f x" by simp + from that[OF h_holo this] show ?thesis . +qed + +subsection\Morera's theorem\ + +lemma Morera_local_triangle_ball: + assumes "\z. z \ S + \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ + (\b c. closed_segment b c \ ball a e + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + shows "f analytic_on S" +proof - + { fix z assume "z \ S" + with assms obtain e a where + "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" + and 0: "\b c. closed_segment b c \ ball a e + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + by fastforce + 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 + qed (simp add: az) + } + then show ?thesis + by (simp add: analytic_on_def) +qed + +lemma Morera_local_triangle: + assumes "\z. z \ S + \ \t. open t \ z \ t \ continuous_on t f \ + (\a b c. convex hull {a,b,c} \ t + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + shows "f analytic_on S" +proof - + { fix z assume "z \ S" + with assms obtain t where + "open t" and z: "z \ t" and contf: "continuous_on t f" + and 0: "\a b c. convex hull {a,b,c} \ t + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + by force + then obtain e where "e>0" and e: "ball z e \ t" + using open_contains_ball by blast + have [simp]: "continuous_on (ball z e) f" using contf + using continuous_on_subset e by blast + have eq0: "\b c. closed_segment b c \ ball z e \ + contour_integral (linepath z b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c z) f = 0" + by (meson 0 z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) + have "\e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ + (\b c. closed_segment b c \ ball a e \ + contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" + using \e > 0\ eq0 by force + } + then show ?thesis + by (simp add: Morera_local_triangle_ball) +qed + +proposition Morera_triangle: + "\continuous_on S f; open S; + \a b c. convex hull {a,b,c} \ S + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0\ + \ f analytic_on S" + using Morera_local_triangle by blast + +subsection\Combining theorems for higher derivatives including Leibniz rule\ + +lemma higher_deriv_linear [simp]: + "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" + by (induction n) auto + +lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" + by (induction n) auto + +lemma higher_deriv_ident [simp]: + "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" + apply (induction n, simp) + apply (metis higher_deriv_linear lambda_one) + done + +lemma higher_deriv_id [simp]: + "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" + by (simp add: id_def) + +lemma has_complex_derivative_funpow_1: + "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" + apply (induction n, auto) + apply (simp add: id_def) + by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) + +lemma higher_deriv_uminus: + assumes "f holomorphic_on S" "open S" and z: "z \ S" + shows "(deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + have "((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 + then show ?case + by (simp add: DERIV_imp_deriv) +qed + +lemma higher_deriv_add: + fixes z::complex + assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" + shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + have "((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 + 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" + 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 + +lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" + by (cases k) simp_all + +lemma higher_deriv_mult: + fixes z::complex + assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" + shows "(deriv ^^ n) (\w. f w * g w) z = + (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + "\n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + have sumeq: "(\i = 0..n. + of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) = + g z * deriv ((deriv ^^ n) f) z + (\i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" + apply (simp add: bb algebra_simps sum.distrib) + apply (subst (4) sum_Suc_reindex) + apply (auto simp: algebra_simps Suc_diff_le intro: sum.cong) + done + have "((deriv ^^ n) (\w. f w * g w) has_field_derivative + (\i = 0..Suc n. (Suc n choose i) * (deriv ^^ i) f z * (deriv ^^ (Suc n - i)) g z)) + (at z)" + apply (rule has_field_derivative_transform_within_open + [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) + 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 + then show ?case + unfolding funpow.simps o_apply + by (simp add: DERIV_imp_deriv) +qed + +lemma higher_deriv_transform_within_open: + fixes z::complex + assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" + and fg: "\w. w \ S \ f w = g w" + shows "(deriv ^^ i) f z = (deriv ^^ i) g z" +using z +by (induction i arbitrary: z) + (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) + +lemma higher_deriv_compose_linear: + fixes z::complex + assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \ S" + and fg: "\w. w \ S \ u * w \ T" + shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have holo0: "f holomorphic_on (*) u ` S" + by (meson fg f holomorphic_on_subset image_subset_iff) + have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S" + by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T) + have holo3: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S" + by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros) + have 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 "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 + 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 + also have "\ = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" + apply (subst complex_derivative_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 + finally show ?case + by simp +qed + +lemma higher_deriv_add_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_add show ?thesis + by (auto simp: analytic_at_two) +qed + +lemma higher_deriv_diff_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_diff show ?thesis + by (auto simp: analytic_at_two) +qed + +lemma higher_deriv_uminus_at: + "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" + using higher_deriv_uminus + by (auto simp: analytic_at) + +lemma higher_deriv_mult_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w * g w) z = + (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_mult show ?thesis + by (auto simp: analytic_at_two) +qed + + +text\ Nonexistence of isolated singularities and a stronger integral formula.\ + +proposition no_isolated_singularity: + fixes z::complex + assumes f: "continuous_on S f" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" + shows "f holomorphic_on S" +proof - + { fix z + assume "z \ S" and cdf: "\x. x \ S - K \ f field_differentiable at x" + have "f field_differentiable at z" + proof (cases "z \ K") + case False then show ?thesis by (blast intro: cdf \z \ S\) + next + case True + with finite_set_avoid [OF K, of z] + obtain d where "d>0" and d: "\x. \x\K; x \ z\ \ d \ dist z x" + by blast + obtain e where "e>0" and e: "ball z e \ S" + using S \z \ S\ by (force simp: open_contains_ball) + have fde: "continuous_on (ball z (min d e)) f" + by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) + have cont: "{a,b,c} \ ball z (min d e) \ continuous_on (convex hull {a, b, c}) f" for a b c + by (simp add: hull_minimal continuous_on_subset [OF fde]) + have fd: "\{a,b,c} \ ball z (min d e); x \ interior (convex hull {a, b, c}) - K\ + \ f field_differentiable at x" for a b c x + by (metis cdf Diff_iff Int_iff ball_min_Int subsetD convex_ball e interior_mono interior_subset subset_hull) + obtain g where "\w. w \ ball z (min d e) \ (g has_field_derivative f w) (at w within ball z (min d e))" + apply (rule contour_integral_convex_primitive + [OF convex_ball fde Cauchy_theorem_triangle_cofinite [OF _ K]]) + using cont fd by auto + then have "f holomorphic_on ball z (min d e)" + by (metis open_ball at_within_open derivative_is_holomorphic) + then show ?thesis + unfolding holomorphic_on_def + by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) + qed + } + with holf S K show ?thesis + by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) +qed + +lemma no_isolated_singularity': + fixes z::complex + assumes f: "\z. z \ K \ (f \ f z) (at z within S)" + and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" + shows "f holomorphic_on S" +proof (rule no_isolated_singularity[OF _ assms(2-)]) + show "continuous_on S f" unfolding continuous_on_def + proof + fix z assume z: "z \ S" + show "(f \ f z) (at z within S)" + proof (cases "z \ K") + case False + from holf have "continuous_on (S - K) f" + by (rule holomorphic_on_imp_continuous_on) + with z False have "(f \ f z) (at z within (S - K))" + by (simp add: continuous_on_def) + also from z K S False have "at z within (S - K) = at z within S" + by (subst (1 2) at_within_open) (auto intro: finite_imp_closed) + finally show "(f \ f z) (at z within S)" . + qed (insert assms z, simp_all) + qed +qed + +proposition Cauchy_integral_formula_convex: + assumes S: "convex S" and K: "finite K" and contf: "continuous_on S f" + and fcd: "(\x. x \ interior S - K \ f field_differentiable at x)" + and z: "z \ interior S" and vpg: "valid_path \" + and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + have *: "\x. x \ interior S \ f field_differentiable at x" + unfolding holomorphic_on_open [symmetric] field_differentiable_def + using no_isolated_singularity [where S = "interior S"] + by (meson K contf continuous_at_imp_continuous_on continuous_on_interior fcd + field_differentiable_at_within field_differentiable_def holomorphic_onI + holomorphic_on_imp_differentiable_at open_interior) + show ?thesis + by (rule Cauchy_integral_formula_weak [OF S finite.emptyI contf]) (use * assms in auto) +qed + +text\ Formula for higher derivatives.\ + +lemma Cauchy_has_contour_integral_higher_derivative_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) + (circlepath z r)" +using w +proof (induction k arbitrary: w) + case 0 then show ?case + using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) +next + case (Suc k) + have [simp]: "r > 0" using w + using ball_eq_empty by fastforce + have f: "continuous_on (path_image (circlepath z r)) f" + by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def less_imp_le) + obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" + using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] + by (auto simp: contour_integrable_on_def) + then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" + by (rule contour_integral_unique) + have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" + by (force simp: field_differentiable_def) + have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = + of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" + by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) + also have "\ = of_nat (Suc k) * X" + by (simp only: con) + finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . + then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" + by (metis deriv_cmult dnf_diff) + then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" + by (simp add: field_simps) + then show ?case + using of_nat_eq_0_iff X by fastforce +qed + +lemma Cauchy_higher_derivative_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" + (is "?thes2") +proof - + have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) + (circlepath z r)" + using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] + by simp + show ?thes1 using * + using contour_integrable_on_def by blast + show ?thes2 + unfolding contour_integral_unique [OF *] by (simp add: field_split_simps) +qed + +corollary Cauchy_contour_integral_circlepath: + assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" + shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" +by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) + +lemma Cauchy_contour_integral_circlepath_2: + assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" + shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" + using Cauchy_contour_integral_circlepath [OF assms, of 1] + by (simp add: power2_eq_square) + + +subsection\A holomorphic function is analytic, i.e. has local power series\ + +theorem holomorphic_power_series: + assumes holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" +proof - + \ \Replacing \<^term>\r\ and the original (weak) premises with stronger ones\ + obtain r where "r > 0" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" + proof + have "cball z ((r + dist w z) / 2) \ ball z r" + using w by (simp add: dist_commute field_sum_of_halves subset_eq) + then show "f holomorphic_on cball z ((r + dist w z) / 2)" + by (rule holomorphic_on_subset [OF holf]) + have "r > 0" + using w by clarsimp (metis dist_norm le_less_trans norm_ge_zero) + then show "0 < (r + dist w z) / 2" + by simp (use zero_le_dist [of w z] in linarith) + qed (use w in \auto simp: dist_commute\) + then have holf: "f holomorphic_on ball z r" + using ball_subset_cball holomorphic_on_subset by blast + have contf: "continuous_on (cball z r) f" + by (simp add: holfc holomorphic_on_imp_continuous_on) + have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" + by (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) (simp add: \0 < r\) + obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" + by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) + obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" + and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" + proof + show "\u. cmod (u - z) = r \ r - dist z w \ cmod (u - w)" + by (metis add_diff_eq diff_add_cancel dist_norm norm_diff_ineq) + qed (use w in \auto simp: dist_norm norm_minus_commute\) + have ul: "uniform_limit (sphere z r) (\n x. (\kx. f x / (x - w)) sequentially" + unfolding uniform_limit_iff dist_norm + proof clarify + fix e::real + assume "0 < e" + have rr: "0 \ (r - k) / r" "(r - k) / r < 1" using k by auto + obtain n where n: "((r - k) / r) ^ n < e / B * k" + using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force + have "norm ((\k N" and r: "r = dist z u" for N u + proof - + have N: "((r - k) / r) ^ N < e / B * k" + apply (rule le_less_trans [OF power_decreasing n]) + using \n \ N\ k by auto + have u [simp]: "(u \ z) \ (u \ w)" + using \0 < r\ r w by auto + have wzu_not1: "(w - z) / (u - z) \ 1" + by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) + have "norm ((\kk = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)" + using \0 < B\ + apply (auto simp: geometric_sum [OF wzu_not1]) + apply (simp add: field_simps norm_mult [symmetric]) + done + also have "\ = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" + using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) + also have "\ = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" + by (simp add: algebra_simps) + also have "\ = norm (w - z) ^ N * norm (f u) / r ^ N" + by (simp add: norm_mult norm_power norm_minus_commute) + also have "\ \ (((r - k)/r)^N) * B" + using \0 < r\ w k + 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 + also have "\ < e * k" + using \0 < B\ N by (simp add: divide_simps) + also have "\ \ e * norm (u - w)" + using r kle \0 < e\ by (simp add: dist_commute dist_norm) + finally show ?thesis + by (simp add: field_split_simps norm_divide del: power_Suc) + qed + with \0 < r\ show "\\<^sub>F n in sequentially. \x\sphere z r. + norm ((\k\<^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) + done + have cic: "\u. (\y. \k0 < 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) + 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 + then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) + sums (2 * of_real pi * \ * f w)" + using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) + then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) + sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" + by (rule sums_divide) + then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) + sums f w" + by (simp add: field_simps) + then show ?thesis + by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) +qed + + +subsection\The Liouville theorem and the Fundamental Theorem of Algebra\ + +text\ These weak Liouville versions don't even need the derivative formula.\ + +lemma Liouville_weak_0: + assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" + shows "f z = 0" +proof (rule ccontr) + assume fz: "f z \ 0" + with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] + obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" + by (auto simp: dist_norm) + define R where "R = 1 + \B\ + norm z" + have "R > 0" unfolding R_def + proof - + have "0 \ cmod z + \B\" + by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) + then show "0 < 1 + \B\ + cmod z" + by linarith + qed + have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" + apply (rule Cauchy_integral_circlepath) + using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ + done + have "cmod (x - z) = R \ cmod (f x) * 2 < cmod (f z)" for x + unfolding R_def + by (rule B) (use norm_triangle_ineq4 [of x z] in auto) + with \R > 0\ fz show False + using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] + by (auto simp: less_imp_le norm_mult norm_divide field_split_simps) +qed + +proposition Liouville_weak: + assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" + shows "f z = l" + using Liouville_weak_0 [of "\z. f z - l"] + by (simp add: assms holomorphic_on_diff LIM_zero) + +proposition Liouville_weak_inverse: + assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" + obtains z where "f z = 0" +proof - + { assume f: "\z. f z \ 0" + have 1: "(\x. 1 / f x) holomorphic_on UNIV" + by (simp add: holomorphic_on_divide assms f) + have 2: "((\x. 1 / f x) \ 0) at_infinity" + 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 + have False + using Liouville_weak_0 [OF 1 2] f by simp + } + then show ?thesis + using that by blast +qed + +text\In particular we get the Fundamental Theorem of Algebra.\ + +theorem fundamental_theorem_of_algebra: + fixes a :: "nat \ complex" + assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" + obtains z where "(\i\n. a i * z^i) = 0" +using assms +proof (elim disjE bexE) + assume "a 0 = 0" then show ?thesis + by (auto simp: that [of 0]) +next + fix i + assume i: "i \ {1..n}" and nz: "a i \ 0" + have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" + by (rule holomorphic_intros)+ + show thesis + proof (rule Liouville_weak_inverse [OF 1]) + show "\\<^sub>F x in at_infinity. B \ cmod (\i\n. a i * x ^ i)" for B + using i polyfun_extremal nz by force + qed (use that in auto) +qed + +subsection\Weierstrass convergence theorem\ + +lemma holomorphic_uniform_limit: + assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" + and ulim: "uniform_limit (cball z r) f g F" + and F: "\ trivial_limit F" + obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" +proof (cases r "0::real" rule: linorder_cases) + case less then show ?thesis by (force simp: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) +next + case equal then show ?thesis + by (force simp: holomorphic_on_def intro: that) +next + case greater + have contg: "continuous_on (cball z r) g" + using cont uniform_limit_theorem [OF eventually_mono ulim F] by blast + have "path_image (circlepath z r) \ cball z r" + using \0 < r\ by auto + then have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" + by (intro continuous_intros continuous_on_subset [OF contg]) + have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" + if w: "w \ ball z r" for w + proof - + define d where "d = (r - norm(w - z))" + have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) + have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" + unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) + have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" + 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) + 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 + have g_cint: "(\u. g u/(u - w)) contour_integrable_on circlepath z r" + by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) + have cif_tends_cig: "((\n. contour_integral(circlepath z r) (\u. f n u / (u - w))) \ contour_integral(circlepath z r) (\u. g u/(u - w))) F" + by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) + have f_tends_cig: "((\n. 2 * of_real pi * \ * f n w) \ contour_integral (circlepath z r) (\u. g u / (u - w))) F" + proof (rule Lim_transform_eventually) + show "\\<^sub>F x in F. contour_integral (circlepath z r) (\u. f x u / (u - w)) + = 2 * of_real pi * \ * f x w" + 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 fastforce + then have "((\u. g u / (2 * of_real pi * \ * (u - w))) has_contour_integral g w) (circlepath z r)" + using has_contour_integral_div [where c = "2 * of_real pi * \"] + by (force simp: field_simps) + then show ?thesis + by (simp add: dist_norm) + qed + show ?thesis + using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] + by (fastforce simp add: holomorphic_on_open contg intro: that) +qed + + +text\ Version showing that the limit is the limit of the derivatives.\ + +proposition has_complex_derivative_uniform_limit: + fixes z::complex + assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ + (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" + and ulim: "uniform_limit (cball z r) f g F" + and F: "\ trivial_limit F" and "0 < r" + obtains g' where + "continuous_on (cball z r) g" + "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" +proof - + let ?conint = "contour_integral (circlepath z r)" + have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" + by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] ulim F]; + auto simp: holomorphic_on_open field_differentiable_def)+ + then obtain g' where g': "\x. x \ ball z r \ (g has_field_derivative g' x) (at x)" + using DERIV_deriv_iff_has_field_derivative + by (fastforce simp add: holomorphic_on_open) + then have derg: "\x. x \ ball z r \ deriv g x = g' x" + by (simp add: DERIV_imp_deriv) + have tends_f'n_g': "((\n. f' n w) \ g' w) F" if w: "w \ ball z r" for w + proof - + have eq_f': "?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \)" + if cont_fn: "continuous_on (cball z r) (f n)" + and fnd: "\w. w \ ball z r \ (f n has_field_derivative f' n w) (at w)" for n + proof - + have hol_fn: "f n holomorphic_on ball z r" + using fnd by (force simp: holomorphic_on_open) + have "(f n has_field_derivative 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)) (at w)" + by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) + then have f': "f' n w = 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)" + using DERIV_unique [OF fnd] w by blast + show ?thesis + by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] field_split_simps) + qed + define d where "d = (r - norm(w - z))^2" + have "d > 0" + using w by (simp add: dist_commute dist_norm d_def) + have dle: "d \ cmod ((y - w)\<^sup>2)" if "r = cmod (z - y)" for y + proof - + have "w \ ball z (cmod (z - y))" + using that w by fastforce + then have "cmod (w - z) \ cmod (z - y)" + by (simp add: dist_complex_def norm_minus_commute) + moreover have "cmod (z - y) - cmod (w - z) \ cmod (y - w)" + by (metis diff_add_cancel diff_add_eq_diff_diff_swap norm_minus_commute norm_triangle_ineq2) + ultimately show ?thesis + using that by (simp add: d_def norm_power power_mono) + qed + have 1: "\\<^sub>F n in F. (\x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" + by (force simp: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) + have 2: "uniform_limit (sphere z r) (\n x. f n x / (x - w)\<^sup>2) (\x. g x / (x - w)\<^sup>2) F" + unfolding uniform_limit_iff + proof clarify + fix e::real + assume "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 + qed + have "((\n. contour_integral (circlepath z r) (\x. f n x / (x - w)\<^sup>2)) + \ contour_integral (circlepath z r) ((\x. g x / (x - w)\<^sup>2))) F" + by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \0 < r\]) + then have tendsto_0: "((\n. 1 / (2 * of_real pi * \) * (?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2))) \ 0) F" + using Lim_null by (force intro!: tendsto_mult_right_zero) + have "((\n. f' n w - g' w) \ 0) F" + apply (rule Lim_transform_eventually [OF tendsto_0]) + apply (force simp: divide_simps intro: eq_f' eventually_mono [OF cont]) + done + then show ?thesis using Lim_null by blast + qed + obtain g' where "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" + by (blast intro: tends_f'n_g' g') + then show ?thesis using g + using that by blast +qed + + +subsection\<^marker>\tag unimportant\ \Some more simple/convenient versions for applications\ + +lemma holomorphic_uniform_sequence: + assumes S: "open S" + and hol_fn: "\n. (f n) holomorphic_on S" + and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" + shows "g holomorphic_on S" +proof - + have "\f'. (g has_field_derivative f') (at z)" if "z \ S" for z + proof - + obtain r where "0 < r" and r: "cball z r \ S" + and ul: "uniform_limit (cball z r) f g sequentially" + using ulim_g [OF \z \ S\] by blast + have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ f n holomorphic_on ball z r" + proof (intro eventuallyI conjI) + show "continuous_on (cball z r) (f x)" for x + using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r by blast + show "f x holomorphic_on ball z r" for x + by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) + qed + show ?thesis + apply (rule holomorphic_uniform_limit [OF *]) + using \0 < r\ centre_in_ball ul + apply (auto simp: holomorphic_on_open) + done + qed + with S show ?thesis + by (simp add: holomorphic_on_open) +qed + +lemma has_complex_derivative_uniform_sequence: + fixes S :: "complex set" + assumes S: "open S" + and hfd: "\n x. x \ S \ ((f n) has_field_derivative f' n x) (at x)" + and ulim_g: "\x. x \ S + \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" + shows "\g'. \x \ S. (g has_field_derivative g' x) (at x) \ ((\n. f' n x) \ g' x) sequentially" +proof - + have y: "\y. (g has_field_derivative y) (at z) \ (\n. f' n z) \ y" if "z \ S" for z + proof - + obtain r where "0 < r" and r: "cball z r \ S" + and ul: "uniform_limit (cball z r) f g sequentially" + using ulim_g [OF \z \ S\] by blast + have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ + (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))" + proof (intro eventuallyI conjI ballI) + show "continuous_on (cball z r) (f x)" for x + by (meson S continuous_on_subset hfd holomorphic_on_imp_continuous_on holomorphic_on_open r) + show "w \ ball z r \ (f x has_field_derivative f' x w) (at w)" for w x + using ball_subset_cball hfd r by blast + qed + show ?thesis + by (rule has_complex_derivative_uniform_limit [OF *, of g]) (use \0 < r\ ul in \force+\) + qed + show ?thesis + by (rule bchoice) (blast intro: y) +qed + +subsection\On analytic functions defined by a series\ + +lemma series_and_derivative_comparison: + fixes S :: "complex set" + assumes S: "open S" + and h: "summable h" + and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" + and to_g: "\\<^sub>F n in sequentially. \x\S. norm (f n x) \ h n" + obtains g g' where "\x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" +proof - + obtain g where g: "uniform_limit S (\n x. \id>0. cball x d \ S \ uniform_limit (cball x d) (\n x. \i S" for x + proof - + obtain d where "d>0" and d: "cball x d \ S" + using open_contains_cball [of "S"] \x \ S\ S by blast + show ?thesis + proof (intro conjI exI) + show "uniform_limit (cball x d) (\n x. \id > 0\ d in auto) + qed + have "\x. x \ S \ (\n. \i g x" + by (metis tendsto_uniform_limitI [OF g]) + moreover have "\g'. \x\S. (g has_field_derivative g' x) (at x) \ (\n. \i g' x" + by (rule has_complex_derivative_uniform_sequence [OF S]) (auto intro: * hfd DERIV_sum)+ + ultimately show ?thesis + by (metis sums_def that) +qed + +text\A version where we only have local uniform/comparative convergence.\ + +lemma series_and_derivative_comparison_local: + fixes S :: "complex set" + assumes S: "open S" + and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" + and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ (\\<^sub>F n in sequentially. \y\ball x d \ S. norm (f n y) \ h n)" + shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" +proof - + have "\y. (\n. f n z) sums (\n. f n z) \ (\n. f' n z) sums y \ ((\x. \n. f n x) has_field_derivative y) (at z)" + if "z \ S" for z + proof - + obtain d h where "0 < d" "summable h" and le_h: "\\<^sub>F n in sequentially. \y\ball z d \ S. norm (f n y) \ h n" + using to_g \z \ S\ by meson + then obtain r where "r>0" and r: "ball z r \ ball z d \ S" using \z \ S\ S + by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) + have 1: "open (ball z d \ S)" + by (simp add: open_Int S) + have 2: "\n x. x \ ball z d \ S \ (f n has_field_derivative f' n x) (at x)" + by (auto simp: hfd) + obtain g g' where gg': "\x \ ball z d \ S. ((\n. f n x) sums g x) \ + ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" + by (auto intro: le_h series_and_derivative_comparison [OF 1 \summable h\ hfd]) + then have "(\n. f' n z) sums g' z" + by (meson \0 < r\ centre_in_ball contra_subsetD r) + moreover have "(\n. f n z) sums (\n. f n z)" + using summable_sums centre_in_ball \0 < d\ \summable h\ le_h + by (metis (full_types) Int_iff gg' summable_def that) + moreover have "((\x. \n. f n x) has_field_derivative g' z) (at z)" + proof (rule has_field_derivative_transform_within) + show "\x. dist x z < r \ g x = (\n. f n x)" + by (metis subsetD dist_commute gg' mem_ball r sums_unique) + qed (use \0 < r\ gg' \z \ S\ \0 < d\ in auto) + ultimately show ?thesis by auto + qed + then show ?thesis + by (rule_tac x="\x. suminf (\n. f n x)" in exI) meson +qed + + +text\Sometimes convenient to compare with a complex series of positive reals. (?)\ + +lemma series_and_derivative_comparison_complex: + fixes S :: "complex set" + assumes S: "open S" + and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" + and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" + shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" +apply (rule series_and_derivative_comparison_local [OF S hfd], assumption) +apply (rule ex_forward [OF to_g], assumption) +apply (erule exE) +apply (rule_tac x="Re \ h" in exI) +apply (force simp: summable_Re o_def nonneg_Reals_cmod_eq_Re image_subset_iff) +done + +text\Sometimes convenient to compare with a complex series of positive reals. (?)\ +lemma series_differentiable_comparison_complex: + fixes S :: "complex set" + assumes S: "open S" + and hfd: "\n x. x \ S \ f n field_differentiable (at x)" + and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" + obtains g where "\x \ S. ((\n. f n x) sums g x) \ g field_differentiable (at x)" +proof - + have hfd': "\n x. x \ S \ (f n has_field_derivative deriv (f n) x) (at x)" + using hfd field_differentiable_derivI by blast + have "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. deriv (f n) x) sums g' x) \ (g has_field_derivative g' x) (at x)" + by (metis series_and_derivative_comparison_complex [OF S hfd' to_g]) + then show ?thesis + using field_differentiable_def that by blast +qed + +text\In particular, a power series is analytic inside circle of convergence.\ + +lemma power_series_and_derivative_0: + fixes a :: "nat \ complex" and r::real + assumes "summable (\n. a n * r^n)" + shows "\g g'. \z. cmod z < r \ + ((\n. a n * z^n) sums g z) \ ((\n. of_nat n * a n * z^(n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" +proof (cases "0 < r") + case True + have der: "\n z. ((\x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)" + by (rule derivative_eq_intros | simp)+ + have y_le: "\cmod (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 "summable (\n. a n * complex_of_real r ^ n)" + using assms \r > 0\ by simp + moreover have "\z. cmod z < r \ cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)" + using \r > 0\ + by (simp flip: of_real_add) + ultimately have sum: "\z. cmod z < r \ summable (\n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)" + by (rule power_series_conv_imp_absconv_weak) + have "\g g'. \z \ ball 0 r. (\n. (a n) * z ^ n) sums g z \ + (\n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \ (g has_field_derivative g' z) (at z)" + apply (rule series_and_derivative_comparison_complex [OF open_ball der]) + apply (rule_tac x="(r - norm z)/2" in exI) + apply (rule_tac x="\n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI) + using \r > 0\ + apply (auto simp: sum eventually_sequentially norm_mult norm_power dist_norm intro!: mult_left_mono power_mono y_le) + done + then show ?thesis + by (simp add: ball_def) +next + case False then show ?thesis + apply (simp add: not_less) + using less_le_trans norm_not_less_zero by blast +qed + +proposition\<^marker>\tag unimportant\ power_series_and_derivative: + fixes a :: "nat \ complex" and r::real + assumes "summable (\n. a n * r^n)" + obtains g g' where "\z \ ball w r. + ((\n. a n * (z - w) ^ n) sums g z) \ ((\n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \ + (g has_field_derivative g' z) (at z)" + using power_series_and_derivative_0 [OF assms] + apply clarify + apply (rule_tac g="(\z. g(z - w))" in that) + using DERIV_shift [where z="-w"] + apply (auto simp: norm_minus_commute Ball_def dist_norm) + done + +proposition\<^marker>\tag unimportant\ power_series_holomorphic: + assumes "\w. w \ ball z r \ ((\n. a n*(w - z)^n) sums f w)" + shows "f holomorphic_on ball z r" +proof - + have "\f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w + proof - + have inb: "z + complex_of_real ((dist z w + r) / 2) \ ball z r" + proof - + have wz: "cmod (w - z) < r" using w + by (auto simp: field_split_simps dist_norm norm_minus_commute) + then have "0 \ r" + by (meson less_eq_real_def norm_ge_zero order_trans) + show ?thesis + using w by (simp add: dist_norm \0\r\ flip: of_real_add) + qed + have sum: "summable (\n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))" + using assms [OF inb] by (force simp: summable_def dist_norm) + obtain g g' where gg': "\u. u \ ball z ((cmod (z - w) + r) / 2) \ + (\n. a n * (u - z) ^ n) sums g u \ + (\n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \ (g has_field_derivative g' u) (at u)" + by (rule power_series_and_derivative [OF sum, of z]) fastforce + have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u + proof - + have less: "cmod (z - u) * 2 < cmod (z - w) + r" + using that dist_triangle2 [of z u w] + by (simp add: dist_norm [symmetric] algebra_simps) + 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 + qed + have "(f has_field_derivative g' w) (at w)" + by (rule has_field_derivative_transform_within [where d="(r - norm(z - w))/2"]) + (use w gg' [of w] in \(force simp: dist_norm)+\) + then show ?thesis .. + qed + then show ?thesis by (simp add: holomorphic_on_open) +qed + +corollary holomorphic_iff_power_series: + "f holomorphic_on ball z r \ + (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" + apply (intro iffI ballI holomorphic_power_series, assumption+) + apply (force intro: power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) + done + +lemma power_series_analytic: + "(\w. w \ ball z r \ (\n. a n*(w - z)^n) sums f w) \ f analytic_on ball z r" + by (force simp: analytic_on_open intro!: power_series_holomorphic) + +lemma analytic_iff_power_series: + "f analytic_on ball z r \ + (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" + by (simp add: analytic_on_open holomorphic_iff_power_series) + +subsection\<^marker>\tag unimportant\ \Equality between holomorphic functions, on open ball then connected set\ + +lemma holomorphic_fun_eq_on_ball: + "\f holomorphic_on ball z r; g holomorphic_on ball z r; + w \ ball z r; + \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ + \ f w = g w" + apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) + apply (auto simp: holomorphic_iff_power_series) + done + +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 + +lemma holomorphic_fun_eq_0_on_connected: + assumes holf: "f holomorphic_on S" and "open S" + and cons: "connected S" + and der: "\n. (deriv ^^ n) f z = 0" + and "z \ S" "w \ S" + shows "f w = 0" +proof - + have *: "ball x e \ (\n. {w \ S. (deriv ^^ n) f w = 0})" + if "\u. (deriv ^^ u) f x = 0" "ball x e \ S" for x e + proof - + have "\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 (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" + apply (rule open_subset, force) + 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})" + 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" + using \e>0\ e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) + 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 +qed + +lemma holomorphic_fun_eq_on_connected: + assumes "f holomorphic_on S" "g holomorphic_on S" and "open S" "connected S" + and "\n. (deriv ^^ n) f z = (deriv ^^ n) g z" + and "z \ S" "w \ S" + shows "f w = g w" +proof (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" S z, simplified]) + show "(\x. f x - g x) holomorphic_on S" + by (intro assms holomorphic_intros) + show "\n. (deriv ^^ n) (\x. f x - g x) z = 0" + using assms higher_deriv_diff by auto +qed (use assms in auto) + +lemma holomorphic_fun_eq_const_on_connected: + assumes holf: "f holomorphic_on S" and "open S" + and cons: "connected S" + and der: "\n. 0 < n \ (deriv ^^ n) f z = 0" + and "z \ S" "w \ S" + shows "f w = f z" +proof (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" S z, simplified]) + show "(\w. f w - f z) holomorphic_on S" + by (intro assms holomorphic_intros) + show "\n. (deriv ^^ n) (\w. f w - f z) z = 0" + by (subst higher_deriv_diff) (use assms in \auto intro: holomorphic_intros\) +qed (use assms in auto) + +subsection\<^marker>\tag unimportant\ \Some basic lemmas about poles/singularities\ + +lemma pole_lemma: + assumes holf: "f holomorphic_on S" and a: "a \ interior S" + shows "(\z. if z = a then deriv f a + else (f z - f a) / (z - a)) holomorphic_on S" (is "?F holomorphic_on S") +proof - + have F1: "?F field_differentiable (at u within S)" if "u \ S" "u \ a" for u + proof - + have fcd: "f field_differentiable at u within S" + using holf holomorphic_on_def by (simp add: \u \ S\) + have cd: "(\z. (f z - f a) / (z - a)) field_differentiable at u within S" + by (rule fcd derivative_intros | simp add: that)+ + have "0 < dist a u" using that dist_nz by blast + then show ?thesis + by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \u \ S\) + qed + have F2: "?F field_differentiable at a" if "0 < e" "ball a e \ S" for e + proof - + have holfb: "f holomorphic_on ball a e" + by (rule holomorphic_on_subset [OF holf \ball a e \ S\]) + have 2: "?F holomorphic_on ball a e - {a}" + 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 have "f field_differentiable at a" + using holfb \0 < e\ holomorphic_on_imp_differentiable_at by auto + with True show ?thesis + by (auto simp: continuous_at has_field_derivative_iff simp flip: DERIV_deriv_iff_field_differentiable + elim: rev_iffD1 [OF _ LIM_equal]) + next + case False with 2 that show ?thesis + by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) + qed + then have 1: "continuous_on (ball a e) ?F" + by (clarsimp simp: continuous_on_eq_continuous_at) + have "?F holomorphic_on ball a e" + by (auto intro: no_isolated_singularity [OF 1 2]) + with that show ?thesis + by (simp add: holomorphic_on_open field_differentiable_def [symmetric] + field_differentiable_at_within) + qed + show ?thesis + proof + fix x assume "x \ S" show "?F field_differentiable at x within S" + proof (cases "x=a") + case True then show ?thesis + using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) + next + case False with F1 \x \ S\ + show ?thesis by blast + qed + qed +qed + +lemma pole_theorem: + assumes holg: "g holomorphic_on S" and a: "a \ interior S" + and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" + shows "(\z. if z = a then deriv g a + else f z - g a/(z - a)) holomorphic_on S" + using pole_lemma [OF holg a] + by (rule holomorphic_transform) (simp add: eq field_split_simps) + +lemma pole_lemma_open: + assumes "f holomorphic_on S" "open S" + shows "(\z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on S" +proof (cases "a \ S") + case True with assms interior_eq pole_lemma + show ?thesis by fastforce +next + case False with assms show ?thesis + apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) + apply (rule field_differentiable_transform_within [where f = "\z. (f z - f a)/(z - a)" and d = 1]) + apply (rule derivative_intros | force)+ + done +qed + +lemma pole_theorem_open: + assumes holg: "g holomorphic_on S" and S: "open S" + and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" + shows "(\z. if z = a then deriv g a + else f z - g a/(z - a)) holomorphic_on S" + using pole_lemma_open [OF holg S] + by (rule holomorphic_transform) (auto simp: eq divide_simps) + +lemma pole_theorem_0: + assumes holg: "g holomorphic_on S" and a: "a \ interior S" + and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f holomorphic_on S" + using pole_theorem [OF holg a eq] + by (rule holomorphic_transform) (auto simp: eq field_split_simps) + +lemma pole_theorem_open_0: + assumes holg: "g holomorphic_on S" and S: "open S" + and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f holomorphic_on S" + using pole_theorem_open [OF holg S eq] + by (rule holomorphic_transform) (auto simp: eq field_split_simps) + +lemma pole_theorem_analytic: + assumes g: "g analytic_on S" + and eq: "\z. z \ S + \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" + shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" (is "?F analytic_on S") + unfolding analytic_on_def +proof + fix x + assume "x \ S" + with g obtain e where "0 < e" and e: "g holomorphic_on ball x e" + by (auto simp add: analytic_on_def) + obtain d where "0 < d" and d: "\w. w \ ball x d - {a} \ g w = (w - a) * f w" + using \x \ S\ eq by blast + have "?F holomorphic_on ball x (min d e)" + using d e \x \ S\ by (fastforce simp: holomorphic_on_subset subset_ball intro!: pole_theorem_open) + then show "\e>0. ?F holomorphic_on ball x e" + using \0 < d\ \0 < e\ not_le by fastforce +qed + +lemma pole_theorem_analytic_0: + assumes g: "g analytic_on S" + and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f analytic_on S" +proof - + have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" + by auto + show ?thesis + using pole_theorem_analytic [OF g eq] by simp +qed + +lemma pole_theorem_analytic_open_superset: + assumes g: "g analytic_on S" and "S \ T" "open T" + and eq: "\z. z \ T - {a} \ g z = (z - a) * f z" + shows "(\z. if z = a then deriv g a + else f z - g a/(z - a)) analytic_on S" +proof (rule pole_theorem_analytic [OF g]) + fix z + assume "z \ S" + then obtain e where "0 < e" and e: "ball z e \ T" + using assms openE by blast + then show "\d>0. \w\ball z d - {a}. g w = (w - a) * f w" + using eq by auto +qed + +lemma pole_theorem_analytic_open_superset_0: + assumes g: "g analytic_on S" "S \ T" "open T" "\z. z \ T - {a} \ g z = (z - a) * f z" + and [simp]: "f a = deriv g a" "g a = 0" + shows "f analytic_on S" +proof - + have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" + by auto + have "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" + by (rule pole_theorem_analytic_open_superset [OF g]) + then show ?thesis by simp +qed + + +subsection\General, homology form of Cauchy's integral formula\ + +text\Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\ + +lemma contour_integral_continuous_on_linepath_2D: + assumes "open U" and cont_dw: "\w. w \ U \ F w contour_integrable_on (linepath a b)" + and cond_uu: "continuous_on (U \ U) (\(x,y). F x y)" + and abu: "closed_segment a b \ U" + shows "continuous_on U (\w. contour_integral (linepath a b) (F w))" +proof - + have *: "\d>0. \x'\U. dist x' w < d \ + dist (contour_integral (linepath a b) (F x')) + (contour_integral (linepath a b) (F w)) \ \" + if "w \ U" "0 < \" "a \ b" for w \ + proof - + obtain \ where "\>0" and \: "cball w \ \ U" using open_contains_cball \open U\ \w \ U\ by force + let ?TZ = "cball w \ \ closed_segment a b" + have "uniformly_continuous_on ?TZ (\(x,y). F x y)" + proof (rule compact_uniformly_continuous) + show "continuous_on ?TZ (\(x,y). F x y)" + by (rule continuous_on_subset[OF cond_uu]) (use SigmaE \ abu in blast) + show "compact ?TZ" + by (simp add: compact_Times) + qed + then obtain \ where "\>0" + and \: "\x x'. \x\?TZ; x'\?TZ; dist x' x < \\ \ + dist ((\(x,y). F x y) x') ((\(x,y). F x y) x) < \/norm(b - a)" + apply (rule uniformly_continuous_onE [where e = "\/norm(b - a)"]) + using \0 < \\ \a \ b\ by auto + have \: "\norm (w - x1) \ \; x2 \ closed_segment a b; + norm (w - x1') \ \; x2' \ closed_segment a b; norm ((x1', x2') - (x1, x2)) < \\ + \ norm (F x1' x2' - F x1 x2) \ \ / cmod (b - a)" + for x1 x2 x1' x2' + using \ [of "(x1,x2)" "(x1',x2')"] by (force simp: dist_norm) + have le_ee: "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \" + if "x' \ U" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' + proof - + have "(\x. F x' x - F w x) contour_integrable_on linepath a b" + by (simp add: \w \ U\ cont_dw contour_integrable_diff that) + then have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" + apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \]) + using \0 < \\ \0 < \\ that apply (auto simp: norm_minus_commute) + done + 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 + 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 + +text\This version has \<^term>\polynomial_function \\ as an additional assumption.\ +lemma Cauchy_integral_formula_global_weak: + assumes "open U" and holf: "f holomorphic_on U" + and z: "z \ U" and \: "polynomial_function \" + and pasz: "path_image \ \ U - {z}" and loop: "pathfinish \ = pathstart \" + and zero: "\w. w \ U \ winding_number \ w = 0" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + obtain \' where pf\': "polynomial_function \'" and \': "\x. (\ has_vector_derivative (\' x)) (at x)" + using has_vector_derivative_polynomial_function [OF \] by blast + then have "bounded(path_image \')" + by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) + then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" + using bounded_pos by force + define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w + define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" + have "path \" "valid_path \" using \ + by (auto simp: path_polynomial_function valid_path_polynomial_function) + then have ov: "open v" + by (simp add: v_def open_winding_number_levelsets loop) + have uv_Un: "U \ v = UNIV" + using pasz zero by (auto simp: v_def) + have conf: "continuous_on U f" + by (metis holf holomorphic_on_imp_continuous_on) + have hol_d: "(d y) holomorphic_on U" if "y \ U" for y + proof - + have *: "(\c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on U" + by (simp add: holf pole_lemma_open \open U\) + then have "isCont (\x. if x = y then deriv f y else (f x - f y) / (x - y)) y" + using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that \open U\ by fastforce + then have "continuous_on U (d y)" + 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\) + 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: \open U\) + qed + have cint_fxy: "(\x. (f x - f y) / (x - y)) contour_integrable_on \" if "y \ path_image \" for y + proof (rule contour_integrable_holomorphic_simple [where S = "U-{y}"]) + show "(\x. (f x - f y) / (x - y)) holomorphic_on U - {y}" + by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) + show "path_image \ \ U - {y}" + using pasz that by blast + qed (auto simp: \open U\ open_delete \valid_path \\) + define h where + "h z = (if z \ U then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z + have U: "((d z) has_contour_integral h z) \" if "z \ U" for z + proof - + have "d z holomorphic_on U" + by (simp add: hol_d that) + with that show ?thesis + apply (simp add: h_def) + 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" + using v_def z by auto + then have "((\x. 1 / (x - z)) has_contour_integral 0) \" + using z v_def has_contour_integral_winding_number [OF \valid_path \\] by fastforce + then have "((\x. f z * (1 / (x - z))) has_contour_integral 0) \" + using has_contour_integral_lmul by fastforce + then have "((\x. f z / (x - z)) has_contour_integral 0) \" + by (simp add: field_split_simps) + moreover have "((\x. (f x - f z) / (x - z)) has_contour_integral contour_integral \ (d z)) \" + using z + apply (auto simp: v_def) + apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) + done + ultimately have *: "((\x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \ (d z))) \" + by (rule has_contour_integral_add) + have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (d z)) \" + if "z \ U" + using * by (auto simp: divide_simps has_contour_integral_eq) + moreover have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (\w. f w / (w - z))) \" + if "z \ U" + 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 + ultimately show ?thesis + using z by (simp add: h_def) + qed + have znot: "z \ path_image \" + using pasz by blast + obtain d0 where "d0>0" and d0: "\x y. x \ path_image \ \ y \ - U \ d0 \ dist x y" + using separate_compact_closed [of "path_image \" "-U"] pasz \open U\ + 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" + apply (rule that [of "d0/2"]) + using \0 < d0\ + apply (auto simp: dist_norm dest: d0) + done + 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\ + 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 + obtain L where "L>0" + and L: "\f B. \f holomorphic_on interior T; \z. z\interior T \ cmod (f z) \ B\ \ + cmod (contour_integral \ f) \ L * B" + using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] + by blast + have "bounded(f ` T)" + by (meson \compact T\ compact_continuous_image compact_imp_bounded conf continuous_on_subset T) + then obtain D where "D>0" and D: "\x. x \ T \ norm (f x) \ D" + by (auto simp: bounded_pos) + obtain C where "C>0" and C: "\x. x \ T \ norm x \ C" + using \compact T\ bounded_pos compact_imp_bounded by force + have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y + proof - + have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp + with le have ybig: "norm y > C" by force + with C have "y \ T" by force + then have ynot: "y \ path_image \" + using subt interior_subset by blast + have [simp]: "winding_number \ y = 0" + 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 + 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 + have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior T" for z + proof - + have "D * L / e + cmod z \ cmod y" + using le C [of z] z using interior_subset by force + then have DL2: "D * L / e \ cmod (z - y)" + using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) + have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" + by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) + also have "\ \ D * (e / L / D)" + 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 + finally show ?thesis . + qed + have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" + by (simp add: dist_norm) + also have "\ \ L * (D * (e / L / D))" + by (rule L [OF holint leD]) + also have "\ = e" + using \L>0\ \0 < D\ by auto + finally show ?thesis . + qed + then have "(h \ 0) at_infinity" + by (meson Lim_at_infinityI) + moreover have "h holomorphic_on UNIV" + proof - + have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" + if "x \ U" "z \ U" "x \ z" for x z + using that conf + apply (simp add: split_def continuous_on_eq_continuous_at \open U\) + apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ + done + have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" + by (rule continuous_intros)+ + have open_uu_Id: "open (U \ U - Id)" + 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 + 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 + proof (rule Lim_withinI) + fix e::real assume "0 < e" + obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" + using \0 < e\ continuous_within_E [OF con_derf [OF \x \ U\]] + by (metis UNIV_I dist_norm) + obtain k2 where "k2>0" and k2: "ball x k2 \ U" + by (blast intro: openE [OF \open U\] \x \ U\) + have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" + if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" + for x' z' + proof - + have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w + apply (drule segment_furthest_le [where y=x]) + by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) + have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w + by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) + have f_has_der: "\x. x \ U \ (f has_field_derivative deriv f x) (at x within U)" + by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \open U\) + have "closed_segment x' z' \ U" + by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) + then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" + using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp + then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" + by (rule has_contour_integral_div) + have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" + apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) + using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] + \e > 0\ \z' \ x'\ + apply (auto simp: norm_divide divide_simps derf_le) + done + also have "\ \ e" using \0 < e\ by simp + finally show ?thesis . + qed + show "\d>0. \xa\U \ U. + 0 < dist xa (x, x) \ dist xa (x, x) < d \ + dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" + apply (rule_tac x="min k1 k2" in exI) + using \k1>0\ \k2>0\ \e>0\ + apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) + 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" + 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)" + apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') + apply (simp add: tendsto_within_open_NO_MATCH open_Times \open U\, clarify) + apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) + using con_ff + apply (auto simp: continuous_within) + done + have hol_dw: "(\z. d z w) holomorphic_on U" if "w \ U" for w + proof - + have "continuous_on U ((\(x,y). d x y) \ (\z. (w,z)))" + by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ + then have *: "continuous_on U (\z. if w = z then deriv f z else (f w - f z) / (w - z))" + by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) + have **: "\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 + 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 + 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 {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) \ \)" + proof (rule continuous_on_compose) + show "continuous_on {0..1} \" + using \path \\ path_def by blast + show "continuous_on (\ ` {0..1}) (\w. contour_integral (linepath a b) (\z. d z w))" + using pasz unfolding path_image_def + by (auto intro!: continuous_on_subset [OF cont_cint_d]) + qed + have 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 \']) + 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) + finally have cint_h_eq: + "contour_integral (linepath a b) h = + contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . + note cint_cint cint_h_eq + } note cint_h = this + have conthu: "continuous_on U h" + proof (simp add: continuous_on_sequentially, clarify) + fix a x + assume x: "x \ U" and au: "\n. a n \ U" and ax: "a \ x" + then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" + by (meson U contour_integrable_on_def eventuallyI) + obtain dd where "dd>0" and dd: "cball x dd \ U" using open_contains_cball \open U\ x by force + have A2: "uniform_limit (path_image \) (\n. d (a n)) (d x) sequentially" + unfolding uniform_limit_iff dist_norm + proof clarify + fix ee::real + assume "0 < ee" + show "\\<^sub>F n in sequentially. \\\path_image \. cmod (d (a n) \ - d x \) < ee" + proof - + let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" + have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" + 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 + then obtain kk where "kk>0" + and kk: "\x x'. \x \ ?ddpa; x' \ ?ddpa; dist x' x < kk\ \ + dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" + by (rule uniformly_continuous_onE [where e = ee]) (use \0 < ee\ in auto) + have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" + for w z + using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm) + 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 + qed + qed + have "(\n. contour_integral \ (d (a n))) \ contour_integral \ (d x)" + by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \valid_path \\) + then have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" + by (simp add: h_def x) + then show "(h \ a) \ h x" + by (simp add: h_def x au o_def) + qed + show ?thesis + proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) + fix z0 + consider "z0 \ v" | "z0 \ U" using uv_Un by blast + then show "h field_differentiable at z0" + proof cases + assume "z0 \ v" then show ?thesis + using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ + by (auto simp: field_differentiable_def v_def) + next + assume "z0 \ U" then + obtain e where "e>0" and e: "ball z0 e \ U" by (blast intro: openE [OF \open U\]) + have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" + if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c + proof - + have *: "\x1 x2 z. z \ U \ closed_segment x1 x2 \ U \ (\w. d w z) contour_integrable_on linepath x1 x2" + using hol_dw holomorphic_on_imp_continuous_on \open U\ + by (auto intro!: contour_integrable_holomorphic_simple) + have abc: "closed_segment a b \ U" "closed_segment b c \ U" "closed_segment c a \ U" + using that e segments_subset_convex_hull by fastforce+ + have eq0: "\w. w \ U \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" + apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) + apply (rule holomorphic_on_subset [OF hol_dw]) + using e abc_subset by auto + have "contour_integral \ + (\x. contour_integral (linepath a b) (\z. d z x) + + (contour_integral (linepath b c) (\z. d z x) + + contour_integral (linepath c a) (\z. d z x))) = 0" + apply (rule contour_integral_eq_0) + using abc pasz U + apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ + done + then show ?thesis + by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) + qed + show ?thesis + using e \e > 0\ + by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic + Morera_triangle continuous_on_subset [OF conthu] *) + qed + qed + qed + ultimately have [simp]: "h z = 0" for z + by (meson Liouville_weak) + have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" + by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) + then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" + by (metis mult.commute has_contour_integral_lmul) + then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" + by (simp add: field_split_simps) + moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" + using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) + show ?thesis + using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) +qed + +theorem Cauchy_integral_formula_global: + assumes S: "open S" and holf: "f holomorphic_on S" + and z: "z \ S" and vpg: "valid_path \" + and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" + and zero: "\w. w \ S \ winding_number \ w = 0" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + have "path \" using vpg by (blast intro: valid_path_imp_path) + have hols: "(\w. f w / (w - z)) holomorphic_on S - {z}" "(\w. 1 / (w - z)) holomorphic_on S - {z}" + by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ + then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" + by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete S vpg pasz) + obtain d where "d>0" + and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; + pathstart h = pathstart g \ pathfinish h = pathfinish g\ + \ path_image h \ S - {z} \ (\f. f holomorphic_on S - {z} \ contour_integral h f = contour_integral g f)" + using contour_integral_nearby_ends [OF _ \path \\ pasz] S by (simp add: open_Diff) metis + obtain p where polyp: "polynomial_function p" + and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" + using path_approx_polynomial_function [OF \path \\ \d > 0\] by blast + then have ploop: "pathfinish p = pathstart p" using loop by auto + have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast + have [simp]: "z \ path_image \" using pasz by blast + have paps: "path_image p \ S - {z}" and cint_eq: "(\f. f holomorphic_on S - {z} \ contour_integral p f = contour_integral \ f)" + using pf ps led d [OF vpg vpp] \d > 0\ by auto + have wn_eq: "winding_number p z = winding_number \ z" + using vpp paps + by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) + have "winding_number p w = winding_number \ w" if "w \ S" for w + proof - + have hol: "(\v. 1 / (v - w)) holomorphic_on S - {z}" + using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) + have "w \ path_image p" "w \ path_image \" using paps pasz that by auto + then show ?thesis + using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) + qed + then have wn0: "\w. w \ S \ winding_number p w = 0" + by (simp add: zero) + show ?thesis + using Cauchy_integral_formula_global_weak [OF S holf z polyp paps ploop wn0] hols + by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) +qed + +subsection \Generalised Cauchy's integral theorem\ + +theorem Cauchy_theorem_global: + assumes S: "open S" and holf: "f holomorphic_on S" + and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" + and pas: "path_image \ \ S" + and zero: "\w. w \ S \ winding_number \ w = 0" + shows "(f has_contour_integral 0) \" +proof - + obtain z where "z \ S" and znot: "z \ path_image \" + proof - + have "compact (path_image \)" + using compact_valid_path_image vpg by blast + then have "path_image \ \ S" + by (metis (no_types) compact_open path_image_nonempty S) + with pas show ?thesis by (blast intro: that) + qed + then have pasz: "path_image \ \ S - {z}" using pas by blast + have hol: "(\w. (w - z) * f w) holomorphic_on S" + by (rule holomorphic_intros holf)+ + show ?thesis + using Cauchy_integral_formula_global [OF S hol \z \ S\ vpg pasz loop zero] + by (auto simp: znot elim!: has_contour_integral_eq) +qed + +corollary Cauchy_theorem_global_outside: + assumes "open S" "f holomorphic_on S" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ S" + "\w. w \ S \ w \ outside(path_image \)" + shows "(f has_contour_integral 0) \" +by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) + +lemma Cauchy_theorem_simply_connected: + assumes "open S" "simply_connected S" "f holomorphic_on S" "valid_path g" + "path_image g \ S" "pathfinish g = pathstart g" + shows "(f has_contour_integral 0) g" +using assms +apply (simp add: simply_connected_eq_contractible_path) +apply (auto intro!: Cauchy_theorem_null_homotopic [where a = "pathstart g"] + homotopic_paths_imp_homotopic_loops) +using valid_path_imp_path by blast + +end \ No newline at end of file diff --git a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy @@ -1,7847 +1,1695 @@ -section \Complex Path Integrals and Cauchy's Integral Theorem\ +section \Cauchy's Integral Theorem\ -text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ +text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ theory Cauchy_Integral_Theorem imports - Complex_Transcendental - Henstock_Kurzweil_Integration + Contour_Integration Weierstrass_Theorems Retracts begin +subsection \Misc\ + +(*TODO: move. Not used in HOL/Analysis.*) lemma leibniz_rule_holomorphic: fixes f::"complex \ 'b::euclidean_space \ complex" assumes "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes "\x. x \ U \ (f x) integrable_on cbox a b" assumes "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes "convex U" shows "(\x. integral (cbox a b) (f x)) holomorphic_on U" using leibniz_rule_field_differentiable[OF assms(1-3) _ assms(4)] by (auto simp: holomorphic_on_def) +(*TODO: move. Not used in HOL/Analysis.*) lemma Ln_measurable [measurable]: "Ln \ measurable borel borel" proof - have *: "Ln (-of_real x) = of_real (ln x) + \ * pi" if "x > 0" for x using that by (subst Ln_minus) (auto simp: Ln_of_real) have **: "Ln (of_real x) = of_real (ln (-x)) + \ * pi" if "x < 0" for x using *[of "-x"] that by simp have cont: "(\x. indicat_real (- \\<^sub>\\<^sub>0) x *\<^sub>R Ln x) \ borel_measurable borel" by (intro borel_measurable_continuous_on_indicator continuous_intros) auto have "(\x. if x \ \\<^sub>\\<^sub>0 then ln (-Re x) + \ * pi else indicator (-\\<^sub>\\<^sub>0) x *\<^sub>R Ln x) \ borel \\<^sub>M borel" (is "?f \ _") by (rule measurable_If_set[OF _ cont]) auto hence "(\x. if x = 0 then Ln 0 else ?f x) \ borel \\<^sub>M borel" by measurable also have "(\x. if x = 0 then Ln 0 else ?f x) = Ln" by (auto simp: fun_eq_iff ** nonpos_Reals_def) finally show ?thesis . qed +(*TODO: move. Not used in HOL/Analysis.*) lemma powr_complex_measurable [measurable]: assumes [measurable]: "f \ measurable M borel" "g \ measurable M borel" shows "(\x. f x powr g x :: complex) \ measurable M borel" using assms by (simp add: powr_def) -subsection\<^marker>\tag unimportant\ \Homeomorphisms of arc images\ - -lemma homeomorphism_arc: - fixes g :: "real \ 'a::t2_space" - assumes "arc g" - obtains h where "homeomorphism {0..1} (path_image g) g h" -using assms by (force simp: arc_def homeomorphism_compact path_def path_image_def) - -lemma homeomorphic_arc_image_interval: - fixes g :: "real \ 'a::t2_space" and a::real - assumes "arc g" "a < b" - shows "(path_image g) homeomorphic {a..b}" -proof - - have "(path_image g) homeomorphic {0..1::real}" - by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc) - also have "\ homeomorphic {a..b}" - using assms by (force intro: homeomorphic_closed_intervals_real) - finally show ?thesis . -qed - -lemma homeomorphic_arc_images: - fixes g :: "real \ 'a::t2_space" and h :: "real \ 'b::t2_space" - assumes "arc g" "arc h" - shows "(path_image g) homeomorphic (path_image h)" -proof - - have "(path_image g) homeomorphic {0..1::real}" - by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc) - also have "\ homeomorphic (path_image h)" - by (meson assms homeomorphic_def homeomorphism_arc) - finally show ?thesis . -qed - -lemma path_connected_arc_complement: - fixes \ :: "real \ 'a::euclidean_space" - assumes "arc \" "2 \ DIM('a)" - shows "path_connected(- path_image \)" -proof - - have "path_image \ homeomorphic {0..1::real}" - by (simp add: assms homeomorphic_arc_image_interval) - then - show ?thesis - apply (rule path_connected_complement_homeomorphic_convex_compact) - apply (auto simp: assms) - done -qed - -lemma connected_arc_complement: - fixes \ :: "real \ 'a::euclidean_space" - assumes "arc \" "2 \ DIM('a)" - shows "connected(- path_image \)" - by (simp add: assms path_connected_arc_complement path_connected_imp_connected) - -lemma inside_arc_empty: - fixes \ :: "real \ 'a::euclidean_space" - assumes "arc \" - shows "inside(path_image \) = {}" -proof (cases "DIM('a) = 1") - case True - then show ?thesis - using assms connected_arc_image connected_convex_1_gen inside_convex by blast -next - case False - show ?thesis - proof (rule inside_bounded_complement_connected_empty) - show "connected (- path_image \)" - apply (rule connected_arc_complement [OF assms]) - using False - by (metis DIM_ge_Suc0 One_nat_def Suc_1 not_less_eq_eq order_class.order.antisym) - show "bounded (path_image \)" - by (simp add: assms bounded_arc_image) - qed -qed - -lemma inside_simple_curve_imp_closed: - fixes \ :: "real \ 'a::euclidean_space" - shows "\simple_path \; x \ inside(path_image \)\ \ pathfinish \ = pathstart \" - using arc_simple_path inside_arc_empty by blast - - -subsection\<^marker>\tag unimportant\ \Piecewise differentiable functions\ - -definition piecewise_differentiable_on - (infixr "piecewise'_differentiable'_on" 50) - where "f piecewise_differentiable_on i \ - continuous_on i f \ - (\S. finite S \ (\x \ i - S. f differentiable (at x within i)))" - -lemma piecewise_differentiable_on_imp_continuous_on: - "f piecewise_differentiable_on S \ continuous_on S f" -by (simp add: piecewise_differentiable_on_def) - -lemma piecewise_differentiable_on_subset: - "f piecewise_differentiable_on S \ T \ S \ f piecewise_differentiable_on T" - using continuous_on_subset - unfolding piecewise_differentiable_on_def - apply safe - apply (blast elim: continuous_on_subset) - by (meson Diff_iff differentiable_within_subset subsetCE) - -lemma differentiable_on_imp_piecewise_differentiable: - fixes a:: "'a::{linorder_topology,real_normed_vector}" - shows "f differentiable_on {a..b} \ f piecewise_differentiable_on {a..b}" - apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on) - apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def) - done - -lemma differentiable_imp_piecewise_differentiable: - "(\x. x \ S \ f differentiable (at x within S)) - \ f piecewise_differentiable_on S" -by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def - intro: differentiable_within_subset) - -lemma piecewise_differentiable_const [iff]: "(\x. z) piecewise_differentiable_on S" - by (simp add: differentiable_imp_piecewise_differentiable) - -lemma piecewise_differentiable_compose: - "\f piecewise_differentiable_on S; g piecewise_differentiable_on (f ` S); - \x. finite (S \ f-`{x})\ - \ (g \ f) piecewise_differentiable_on S" - apply (simp add: piecewise_differentiable_on_def, safe) - apply (blast intro: continuous_on_compose2) - apply (rename_tac A B) - apply (rule_tac x="A \ (\x\B. S \ f-`{x})" in exI) - apply (blast intro!: differentiable_chain_within) - done - -lemma piecewise_differentiable_affine: - fixes m::real - assumes "f piecewise_differentiable_on ((\x. m *\<^sub>R x + c) ` S)" - shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_differentiable_on S" -proof (cases "m = 0") - case True - then show ?thesis - unfolding o_def - by (force intro: differentiable_imp_piecewise_differentiable differentiable_const) -next - case False - show ?thesis - apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable]) - apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+ - done -qed - -lemma piecewise_differentiable_cases: - fixes c::real - assumes "f piecewise_differentiable_on {a..c}" - "g piecewise_differentiable_on {c..b}" - "a \ c" "c \ b" "f c = g c" - shows "(\x. if x \ c then f x else g x) piecewise_differentiable_on {a..b}" -proof - - obtain S T where st: "finite S" "finite T" - and fd: "\x. x \ {a..c} - S \ f differentiable at x within {a..c}" - and gd: "\x. x \ {c..b} - T \ g differentiable at x within {c..b}" - using assms - by (auto simp: piecewise_differentiable_on_def) - have finabc: "finite ({a,b,c} \ (S \ T))" - by (metis \finite S\ \finite T\ finite_Un finite_insert finite.emptyI) - have "continuous_on {a..c} f" "continuous_on {c..b} g" - using assms piecewise_differentiable_on_def by auto - then have "continuous_on {a..b} (\x. if x \ c then f x else g x)" - using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], - OF closed_real_atLeastAtMost [of c b], - of f g "\x. x\c"] assms - by (force simp: ivl_disj_un_two_touch) - moreover - { fix x - assume x: "x \ {a..b} - ({a,b,c} \ (S \ T))" - have "(\x. if x \ c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg") - proof (cases x c rule: le_cases) - case le show ?diff_fg - proof (rule differentiable_transform_within [where d = "dist x c"]) - have "f differentiable at x" - using x le fd [of x] at_within_interior [of x "{a..c}"] by simp - then show "f differentiable at x within {a..b}" - by (simp add: differentiable_at_withinI) - qed (use x le st dist_real_def in auto) - next - case ge show ?diff_fg - proof (rule differentiable_transform_within [where d = "dist x c"]) - have "g differentiable at x" - using x ge gd [of x] at_within_interior [of x "{c..b}"] by simp - then show "g differentiable at x within {a..b}" - by (simp add: differentiable_at_withinI) - qed (use x ge st dist_real_def in auto) - qed - } - then have "\S. finite S \ - (\x\{a..b} - S. (\x. if x \ c then f x else g x) differentiable at x within {a..b})" - by (meson finabc) - ultimately show ?thesis - by (simp add: piecewise_differentiable_on_def) -qed - -lemma piecewise_differentiable_neg: - "f piecewise_differentiable_on S \ (\x. -(f x)) piecewise_differentiable_on S" - by (auto simp: piecewise_differentiable_on_def continuous_on_minus) - -lemma piecewise_differentiable_add: - assumes "f piecewise_differentiable_on i" - "g piecewise_differentiable_on i" - shows "(\x. f x + g x) piecewise_differentiable_on i" -proof - - obtain S T where st: "finite S" "finite T" - "\x\i - S. f differentiable at x within i" - "\x\i - T. g differentiable at x within i" - using assms by (auto simp: piecewise_differentiable_on_def) - then have "finite (S \ T) \ (\x\i - (S \ T). (\x. f x + g x) differentiable at x within i)" - by auto - moreover have "continuous_on i f" "continuous_on i g" - using assms piecewise_differentiable_on_def by auto - ultimately show ?thesis - by (auto simp: piecewise_differentiable_on_def continuous_on_add) -qed - -lemma piecewise_differentiable_diff: - "\f piecewise_differentiable_on S; g piecewise_differentiable_on S\ - \ (\x. f x - g x) piecewise_differentiable_on S" - unfolding diff_conv_add_uminus - by (metis piecewise_differentiable_add piecewise_differentiable_neg) - -lemma continuous_on_joinpaths_D1: - "continuous_on {0..1} (g1 +++ g2) \ continuous_on {0..1} g1" - apply (rule continuous_on_eq [of _ "(g1 +++ g2) \ ((*)(inverse 2))"]) - apply (rule continuous_intros | simp)+ - apply (auto elim!: continuous_on_subset simp: joinpaths_def) - done - -lemma continuous_on_joinpaths_D2: - "\continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\ \ continuous_on {0..1} g2" - apply (rule continuous_on_eq [of _ "(g1 +++ g2) \ (\x. inverse 2*x + 1/2)"]) - apply (rule continuous_intros | simp)+ - apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def) - done - -lemma piecewise_differentiable_D1: - assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" - shows "g1 piecewise_differentiable_on {0..1}" -proof - - obtain S where cont: "continuous_on {0..1} g1" and "finite S" - and S: "\x. x \ {0..1} - S \ g1 +++ g2 differentiable at x within {0..1}" - using assms unfolding piecewise_differentiable_on_def - by (blast dest!: continuous_on_joinpaths_D1) - show ?thesis - unfolding piecewise_differentiable_on_def - proof (intro exI conjI ballI cont) - show "finite (insert 1 (((*)2) ` S))" - by (simp add: \finite S\) - show "g1 differentiable at x within {0..1}" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x - proof (rule_tac d="dist (x/2) (1/2)" in differentiable_transform_within) - have "g1 +++ g2 differentiable at (x / 2) within {0..1/2}" - by (rule differentiable_subset [OF S [of "x/2"]] | use that in force)+ - then show "g1 +++ g2 \ (*) (inverse 2) differentiable at x within {0..1}" - using image_affinity_atLeastAtMost_div [of 2 0 "0::real" 1] - by (auto intro: differentiable_chain_within) - qed (use that in \auto simp: joinpaths_def\) - qed -qed - -lemma piecewise_differentiable_D2: - assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" and eq: "pathfinish g1 = pathstart g2" - shows "g2 piecewise_differentiable_on {0..1}" -proof - - have [simp]: "g1 1 = g2 0" - using eq by (simp add: pathfinish_def pathstart_def) - obtain S where cont: "continuous_on {0..1} g2" and "finite S" - and S: "\x. x \ {0..1} - S \ g1 +++ g2 differentiable at x within {0..1}" - using assms unfolding piecewise_differentiable_on_def - by (blast dest!: continuous_on_joinpaths_D2) - show ?thesis - unfolding piecewise_differentiable_on_def - proof (intro exI conjI ballI cont) - show "finite (insert 0 ((\x. 2*x-1)`S))" - by (simp add: \finite S\) - show "g2 differentiable at x within {0..1}" if "x \ {0..1} - insert 0 ((\x. 2*x-1)`S)" for x - proof (rule_tac d="dist ((x+1)/2) (1/2)" in differentiable_transform_within) - have x2: "(x + 1) / 2 \ S" - using that - apply (clarsimp simp: image_iff) - by (metis add.commute add_diff_cancel_left' mult_2 field_sum_of_halves) - have "g1 +++ g2 \ (\x. (x+1) / 2) differentiable at x within {0..1}" - by (rule differentiable_chain_within differentiable_subset [OF S [of "(x+1)/2"]] | use x2 that in force)+ - then show "g1 +++ g2 \ (\x. (x+1) / 2) differentiable at x within {0..1}" - by (auto intro: differentiable_chain_within) - show "(g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" if "x' \ {0..1}" "dist x' x < dist ((x + 1) / 2) (1/2)" for x' - proof - - have [simp]: "(2*x'+2)/2 = x'+1" - by (simp add: field_split_simps) - show ?thesis - using that by (auto simp: joinpaths_def) - qed - qed (use that in \auto simp: joinpaths_def\) - qed -qed - - -subsection\The concept of continuously differentiable\ - -text \ -John Harrison writes as follows: - -``The usual assumption in complex analysis texts is that a path \\\ should be piecewise -continuously differentiable, which ensures that the path integral exists at least for any continuous -f, since all piecewise continuous functions are integrable. However, our notion of validity is -weaker, just piecewise differentiability\ldots{} [namely] continuity plus differentiability except on a -finite set\ldots{} [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to -the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this -can integrate all derivatives.'' - -"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec. -Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165. - -And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably -difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem -asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ - -definition\<^marker>\tag important\ C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" - (infix "C1'_differentiable'_on" 50) - where - "f C1_differentiable_on S \ - (\D. (\x \ S. (f has_vector_derivative (D x)) (at x)) \ continuous_on S D)" - -lemma C1_differentiable_on_eq: - "f C1_differentiable_on S \ - (\x \ S. f differentiable at x) \ continuous_on S (\x. vector_derivative f (at x))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - unfolding C1_differentiable_on_def - by (metis (no_types, lifting) continuous_on_eq differentiableI_vector vector_derivative_at) -next - assume ?rhs - then show ?lhs - using C1_differentiable_on_def vector_derivative_works by fastforce -qed - -lemma C1_differentiable_on_subset: - "f C1_differentiable_on T \ S \ T \ f C1_differentiable_on S" - unfolding C1_differentiable_on_def continuous_on_eq_continuous_within - by (blast intro: continuous_within_subset) - -lemma C1_differentiable_compose: - assumes fg: "f C1_differentiable_on S" "g C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" - shows "(g \ f) C1_differentiable_on S" -proof - - have "\x. x \ S \ g \ f differentiable at x" - by (meson C1_differentiable_on_eq assms differentiable_chain_at imageI) - moreover have "continuous_on S (\x. vector_derivative (g \ f) (at x))" - proof (rule continuous_on_eq [of _ "\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) - show "continuous_on S (\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)))" - using fg - apply (clarsimp simp add: C1_differentiable_on_eq) - apply (rule Limits.continuous_on_scaleR, assumption) - by (metis (mono_tags, lifting) continuous_at_imp_continuous_on continuous_on_compose continuous_on_cong differentiable_imp_continuous_within o_def) - show "\x. x \ S \ vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)) = vector_derivative (g \ f) (at x)" - by (metis (mono_tags, hide_lams) C1_differentiable_on_eq fg imageI vector_derivative_chain_at) - qed - ultimately show ?thesis - by (simp add: C1_differentiable_on_eq) -qed - -lemma C1_diff_imp_diff: "f C1_differentiable_on S \ f differentiable_on S" - by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on) - -lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\x. x) C1_differentiable_on S" - by (auto simp: C1_differentiable_on_eq) - -lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" - by (auto simp: C1_differentiable_on_eq) - -lemma C1_differentiable_on_add [simp, derivative_intros]: - "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x + g x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_minus [simp, derivative_intros]: - "f C1_differentiable_on S \ (\x. - f x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_diff [simp, derivative_intros]: - "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x - g x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_mult [simp, derivative_intros]: - fixes f g :: "real \ 'a :: real_normed_algebra" - shows "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x * g x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq - by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) - -lemma C1_differentiable_on_scaleR [simp, derivative_intros]: - "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x *\<^sub>R g x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq - by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ - - -definition\<^marker>\tag important\ piecewise_C1_differentiable_on - (infixr "piecewise'_C1'_differentiable'_on" 50) - where "f piecewise_C1_differentiable_on i \ - continuous_on i f \ - (\S. finite S \ (f C1_differentiable_on (i - S)))" - -lemma C1_differentiable_imp_piecewise: - "f C1_differentiable_on S \ f piecewise_C1_differentiable_on S" - by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) - -lemma piecewise_C1_imp_differentiable: - "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" - by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def - C1_differentiable_on_def differentiable_def has_vector_derivative_def - intro: has_derivative_at_withinI) - -lemma piecewise_C1_differentiable_compose: - assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" - shows "(g \ f) piecewise_C1_differentiable_on S" -proof - - have "continuous_on S (\x. g (f x))" - by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def) - moreover have "\T. finite T \ g \ f C1_differentiable_on S - T" - proof - - obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S" - using fg by (auto simp: piecewise_C1_differentiable_on_def) - obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S" - using fg by (auto simp: piecewise_C1_differentiable_on_def) - show ?thesis - proof (intro exI conjI) - show "finite (F \ (\x\G. S \ f-`{x}))" - using fin by (auto simp only: Int_Union \finite F\ \finite G\ finite_UN finite_imageI) - show "g \ f C1_differentiable_on S - (F \ (\x\G. S \ f -` {x}))" - apply (rule C1_differentiable_compose) - apply (blast intro: C1_differentiable_on_subset [OF F]) - apply (blast intro: C1_differentiable_on_subset [OF G]) - by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin) - qed - qed - ultimately show ?thesis - by (simp add: piecewise_C1_differentiable_on_def) -qed - -lemma piecewise_C1_differentiable_on_subset: - "f piecewise_C1_differentiable_on S \ T \ S \ f piecewise_C1_differentiable_on T" - by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) - -lemma C1_differentiable_imp_continuous_on: - "f C1_differentiable_on S \ continuous_on S f" - unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within - using differentiable_at_withinI differentiable_imp_continuous_within by blast - -lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" - unfolding C1_differentiable_on_def - by auto - -lemma piecewise_C1_differentiable_affine: - fixes m::real - assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` S)" - shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on S" -proof (cases "m = 0") - case True - then show ?thesis - unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def) -next - case False - have *: "\x. finite (S \ {y. m * y + c = x})" - using False not_finite_existsD by fastforce - show ?thesis - apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) - apply (rule * assms derivative_intros | simp add: False vimage_def)+ - done -qed - -lemma piecewise_C1_differentiable_cases: - fixes c::real - assumes "f piecewise_C1_differentiable_on {a..c}" - "g piecewise_C1_differentiable_on {c..b}" - "a \ c" "c \ b" "f c = g c" - shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" -proof - - obtain S T where st: "f C1_differentiable_on ({a..c} - S)" - "g C1_differentiable_on ({c..b} - T)" - "finite S" "finite T" - using assms - by (force simp: piecewise_C1_differentiable_on_def) - then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" - using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], - OF closed_real_atLeastAtMost [of c b], - of f g "\x. x\c"] assms - by (force simp: ivl_disj_un_two_touch) - { fix x - assume x: "x \ {a..b} - insert c (S \ T)" - have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") - proof (cases x c rule: le_cases) - case le show ?diff_fg - apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) - using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) - next - case ge show ?diff_fg - apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) - using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) - qed - } - then have "(\x \ {a..b} - insert c (S \ T). (\x. if x \ c then f x else g x) differentiable at x)" - by auto - moreover - { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" - and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" - have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - proof - - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)" - if "a < x" "x < c" "x \ S" for x - proof - - have f: "f differentiable at x" - by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that) - show ?thesis - using that - apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) - apply (auto simp: dist_norm vector_derivative_works [symmetric] f) - done - qed - then show ?thesis - by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at) - qed - moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - proof - - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)" - if "c < x" "x < b" "x \ T" for x - proof - - have g: "g differentiable at x" - by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that) - show ?thesis - using that - apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) - apply (auto simp: dist_norm vector_derivative_works [symmetric] g) - done - qed - then show ?thesis - by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at) - qed - ultimately have "continuous_on ({a<.. T)) - (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - by (rule continuous_on_subset [OF continuous_on_open_Un], auto) - } note * = this - have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - using st - by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) - ultimately have "\S. finite S \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - S)" - apply (rule_tac x="{a,b,c} \ S \ T" in exI) - using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) - with cab show ?thesis - by (simp add: piecewise_C1_differentiable_on_def) -qed - -lemma piecewise_C1_differentiable_neg: - "f piecewise_C1_differentiable_on S \ (\x. -(f x)) piecewise_C1_differentiable_on S" - unfolding piecewise_C1_differentiable_on_def - by (auto intro!: continuous_on_minus C1_differentiable_on_minus) - -lemma piecewise_C1_differentiable_add: - assumes "f piecewise_C1_differentiable_on i" - "g piecewise_C1_differentiable_on i" - shows "(\x. f x + g x) piecewise_C1_differentiable_on i" -proof - - obtain S t where st: "finite S" "finite t" - "f C1_differentiable_on (i-S)" - "g C1_differentiable_on (i-t)" - using assms by (auto simp: piecewise_C1_differentiable_on_def) - then have "finite (S \ t) \ (\x. f x + g x) C1_differentiable_on i - (S \ t)" - by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) - moreover have "continuous_on i f" "continuous_on i g" - using assms piecewise_C1_differentiable_on_def by auto - ultimately show ?thesis - by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) -qed - -lemma piecewise_C1_differentiable_diff: - "\f piecewise_C1_differentiable_on S; g piecewise_C1_differentiable_on S\ - \ (\x. f x - g x) piecewise_C1_differentiable_on S" - unfolding diff_conv_add_uminus - by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) - -lemma piecewise_C1_differentiable_D1: - fixes g1 :: "real \ 'a::real_normed_field" - assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" - shows "g1 piecewise_C1_differentiable_on {0..1}" -proof - - obtain S where "finite S" - and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" - and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" - using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have g1D: "g1 differentiable at x" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x - proof (rule differentiable_transform_within) - show "g1 +++ g2 \ (*) (inverse 2) differentiable at x" - using that g12D - apply (simp only: joinpaths_def) - by (rule differentiable_chain_at derivative_intros | force)+ - show "\x'. \dist x' x < dist (x/2) (1/2)\ - \ (g1 +++ g2 \ (*) (inverse 2)) x' = g1 x'" - using that by (auto simp: dist_real_def joinpaths_def) - qed (use that in \auto simp: dist_real_def\) - have [simp]: "vector_derivative (g1 \ (*) 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" - if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x - apply (subst vector_derivative_chain_at) - using that - apply (rule derivative_eq_intros g1D | simp)+ - done - have "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" - using co12 by (rule continuous_on_subset) force - then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 \ (*)2) (at x))" - proof (rule continuous_on_eq [OF _ vector_derivative_at]) - show "(g1 +++ g2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" - if "x \ {0..1/2} - insert (1/2) S" for x - proof (rule has_vector_derivative_transform_within) - show "(g1 \ (*) 2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" - using that - by (force intro: g1D differentiable_chain_at simp: vector_derivative_works [symmetric]) - show "\x'. \dist x' x < dist x (1/2)\ \ (g1 \ (*) 2) x' = (g1 +++ g2) x'" - using that by (auto simp: dist_norm joinpaths_def) - qed (use that in \auto simp: dist_norm\) - qed - have "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) - ((\x. 1/2 * vector_derivative (g1 \ (*)2) (at x)) \ (*)(1/2))" - apply (rule continuous_intros)+ - using coDhalf - apply (simp add: scaleR_conv_of_real image_set_diff image_image) - done - then have con_g1: "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) (\x. vector_derivative g1 (at x))" - by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) - have "continuous_on {0..1} g1" - using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast - with \finite S\ show ?thesis - apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - apply (rule_tac x="insert 1 (((*)2)`S)" in exI) - apply (simp add: g1D con_g1) - done -qed - -lemma piecewise_C1_differentiable_D2: - fixes g2 :: "real \ 'a::real_normed_field" - assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" - shows "g2 piecewise_C1_differentiable_on {0..1}" -proof - - obtain S where "finite S" - and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" - and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" - using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have g2D: "g2 differentiable at x" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x - proof (rule differentiable_transform_within) - show "g1 +++ g2 \ (\x. (x + 1) / 2) differentiable at x" - using g12D that - apply (simp only: joinpaths_def) - apply (drule_tac x= "(x+1) / 2" in bspec, force simp: field_split_simps) - apply (rule differentiable_chain_at derivative_intros | force)+ - done - show "\x'. dist x' x < dist ((x + 1) / 2) (1/2) \ (g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" - using that by (auto simp: dist_real_def joinpaths_def field_simps) - qed (use that in \auto simp: dist_norm\) - have [simp]: "vector_derivative (g2 \ (\x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" - if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x - using that by (auto simp: vector_derivative_chain_at field_split_simps g2D) - have "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" - using co12 by (rule continuous_on_subset) force - then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g2 \ (\x. 2*x-1)) (at x))" - proof (rule continuous_on_eq [OF _ vector_derivative_at]) - show "(g1 +++ g2 has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) - (at x)" - if "x \ {1 / 2..1} - insert (1 / 2) S" for x - proof (rule_tac f="g2 \ (\x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) - show "(g2 \ (\x. 2 * x - 1) has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) - (at x)" - using that by (force intro: g2D differentiable_chain_at simp: vector_derivative_works [symmetric]) - show "\x'. \dist x' x < dist (3 / 4) ((x + 1) / 2)\ \ (g2 \ (\x. 2 * x - 1)) x' = (g1 +++ g2) x'" - using that by (auto simp: dist_norm joinpaths_def add_divide_distrib) - qed (use that in \auto simp: dist_norm\) - qed - have [simp]: "((\x. (x+1) / 2) ` ({0..1} - insert 0 ((\x. 2 * x - 1) ` S))) = ({1/2..1} - insert (1/2) S)" - apply (simp add: image_set_diff inj_on_def image_image) - apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) - done - have "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) - ((\x. 1/2 * vector_derivative (g2 \ (\x. 2*x-1)) (at x)) \ (\x. (x+1)/2))" - by (rule continuous_intros | simp add: coDhalf)+ - then have con_g2: "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) (\x. vector_derivative g2 (at x))" - by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) - have "continuous_on {0..1} g2" - using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast - with \finite S\ show ?thesis - apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - apply (rule_tac x="insert 0 ((\x. 2 * x - 1) ` S)" in exI) - apply (simp add: g2D con_g2) - done -qed - -subsection \Valid paths, and their start and finish\ - -definition\<^marker>\tag important\ valid_path :: "(real \ 'a :: real_normed_vector) \ bool" - where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" - -definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" - where "closed_path g \ g 0 = g 1" - -text\In particular, all results for paths apply\ - -lemma valid_path_imp_path: "valid_path g \ path g" - by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) - -lemma connected_valid_path_image: "valid_path g \ connected(path_image g)" - by (metis connected_path_image valid_path_imp_path) - -lemma compact_valid_path_image: "valid_path g \ compact(path_image g)" - by (metis compact_path_image valid_path_imp_path) - -lemma bounded_valid_path_image: "valid_path g \ bounded(path_image g)" - by (metis bounded_path_image valid_path_imp_path) - -lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" - by (metis closed_path_image valid_path_imp_path) - -lemma valid_path_compose: - assumes "valid_path g" - and der: "\x. x \ path_image g \ f field_differentiable (at x)" - and con: "continuous_on (path_image g) (deriv f)" - shows "valid_path (f \ g)" -proof - - obtain S where "finite S" and g_diff: "g C1_differentiable_on {0..1} - S" - using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto - have "f \ g differentiable at t" when "t\{0..1} - S" for t - proof (rule differentiable_chain_at) - show "g differentiable at t" using \valid_path g\ - by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - S\ that) - next - have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis - then show "f differentiable at (g t)" - using der[THEN field_differentiable_imp_differentiable] by auto - qed - moreover have "continuous_on ({0..1} - S) (\x. vector_derivative (f \ g) (at x))" - proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], - rule continuous_intros) - show "continuous_on ({0..1} - S) (\x. vector_derivative g (at x))" - using g_diff C1_differentiable_on_eq by auto - next - have "continuous_on {0..1} (\x. deriv f (g x))" - using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] - \valid_path g\ piecewise_C1_differentiable_on_def valid_path_def - by blast - then show "continuous_on ({0..1} - S) (\x. deriv f (g x))" - using continuous_on_subset by blast - next - show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" - when "t \ {0..1} - S" for t - proof (rule vector_derivative_chain_at_general[symmetric]) - show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) - next - have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis - then show "f field_differentiable at (g t)" using der by auto - qed - qed - ultimately have "f \ g C1_differentiable_on {0..1} - S" - using C1_differentiable_on_eq by blast - moreover have "path (f \ g)" - apply (rule path_continuous_image[OF valid_path_imp_path[OF \valid_path g\]]) - using der - by (simp add: continuous_at_imp_continuous_on field_differentiable_imp_continuous_at) - ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def - using \finite S\ by auto -qed - -lemma valid_path_uminus_comp[simp]: - fixes g::"real \ 'a ::real_normed_field" - shows "valid_path (uminus \ g) \ valid_path g" -proof - show "valid_path g \ valid_path (uminus \ g)" for g::"real \ 'a" - by (auto intro!: valid_path_compose derivative_intros simp add: deriv_linear[of "-1",simplified]) - then show "valid_path g" when "valid_path (uminus \ g)" - by (metis fun.map_comp group_add_class.minus_comp_minus id_comp that) -qed - -lemma valid_path_offset[simp]: - shows "valid_path (\t. g t - z) \ valid_path g" -proof - show *: "valid_path (g::real\'a) \ valid_path (\t. g t - z)" for g z - unfolding valid_path_def - by (fastforce intro:derivative_intros C1_differentiable_imp_piecewise piecewise_C1_differentiable_diff) - show "valid_path (\t. g t - z) \ valid_path g" - using *[of "\t. g t - z" "-z",simplified] . -qed - - -subsection\Contour Integrals along a path\ - -text\This definition is for complex numbers only, and does not generalise to line integrals in a vector field\ - -text\piecewise differentiable function on [0,1]\ - -definition\<^marker>\tag important\ has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" - (infixr "has'_contour'_integral" 50) - where "(f has_contour_integral i) g \ - ((\x. f(g x) * vector_derivative g (at x within {0..1})) - has_integral i) {0..1}" - -definition\<^marker>\tag important\ contour_integrable_on - (infixr "contour'_integrable'_on" 50) - where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" - -definition\<^marker>\tag important\ contour_integral - where "contour_integral g f \ SOME i. (f has_contour_integral i) g \ \ f contour_integrable_on g \ i=0" - -lemma not_integrable_contour_integral: "\ f contour_integrable_on g \ contour_integral g f = 0" - unfolding contour_integrable_on_def contour_integral_def by blast - -lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" - apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) - using has_integral_unique by blast - -lemma has_contour_integral_eqpath: - "\(f has_contour_integral y) p; f contour_integrable_on \; - contour_integral p f = contour_integral \ f\ - \ (f has_contour_integral y) \" -using contour_integrable_on_def contour_integral_unique by auto - -lemma has_contour_integral_integral: - "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" - by (metis contour_integral_unique contour_integrable_on_def) - -lemma has_contour_integral_unique: - "(f has_contour_integral i) g \ (f has_contour_integral j) g \ i = j" - using has_integral_unique - by (auto simp: has_contour_integral_def) - -lemma has_contour_integral_integrable: "(f has_contour_integral i) g \ f contour_integrable_on g" - using contour_integrable_on_def by blast - -text\Show that we can forget about the localized derivative.\ - -lemma has_integral_localized_vector_derivative: - "((\x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \ - ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}" -proof - - have *: "{a..b} - {a,b} = interior {a..b}" - by (simp add: atLeastAtMost_diff_ends) - show ?thesis - apply (rule has_integral_spike_eq [of "{a,b}"]) - apply (auto simp: at_within_interior [of _ "{a..b}"]) - done -qed - -lemma integrable_on_localized_vector_derivative: - "(\x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \ - (\x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}" - by (simp add: integrable_on_def has_integral_localized_vector_derivative) - -lemma has_contour_integral: - "(f has_contour_integral i) g \ - ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" - by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) - -lemma contour_integrable_on: - "f contour_integrable_on g \ - (\t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" - by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) - -subsection\<^marker>\tag unimportant\ \Reversing a path\ - -lemma valid_path_imp_reverse: - assumes "valid_path g" - shows "valid_path(reversepath g)" -proof - - obtain S where "finite S" and S: "g C1_differentiable_on ({0..1} - S)" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - then have "finite ((-) 1 ` S)" - by auto - moreover have "(reversepath g C1_differentiable_on ({0..1} - (-) 1 ` S))" - unfolding reversepath_def - apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) - using S - by (force simp: finite_vimageI inj_on_def C1_differentiable_on_eq elim!: continuous_on_subset)+ - ultimately show ?thesis using assms - by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) -qed - -lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" - using valid_path_imp_reverse by force - -lemma has_contour_integral_reversepath: - assumes "valid_path g" and f: "(f has_contour_integral i) g" - shows "(f has_contour_integral (-i)) (reversepath g)" -proof - - { fix S x - assume xs: "g C1_differentiable_on ({0..1} - S)" "x \ (-) 1 ` S" "0 \ x" "x \ 1" - have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - - vector_derivative g (at (1 - x) within {0..1})" - proof - - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" - using xs - by (force simp: has_vector_derivative_def C1_differentiable_on_def) - have "(g \ (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" - by (intro vector_diff_chain_within has_vector_derivative_at_within [OF f'] derivative_eq_intros | simp)+ - then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" - by (simp add: o_def) - show ?thesis - using xs - by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) - qed - } note * = this - obtain S where S: "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - have "((\x. - (f (g (1 - x)) * vector_derivative g (at (1 - x) within {0..1}))) has_integral -i) - {0..1}" - using has_integral_affinity01 [where m= "-1" and c=1, OF f [unfolded has_contour_integral_def]] - by (simp add: has_integral_neg) - then show ?thesis - using S - apply (clarsimp simp: reversepath_def has_contour_integral_def) - apply (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) - apply (auto simp: *) - done -qed - -lemma contour_integrable_reversepath: - "valid_path g \ f contour_integrable_on g \ f contour_integrable_on (reversepath g)" - using has_contour_integral_reversepath contour_integrable_on_def by blast - -lemma contour_integrable_reversepath_eq: - "valid_path g \ (f contour_integrable_on (reversepath g) \ f contour_integrable_on g)" - using contour_integrable_reversepath valid_path_reversepath by fastforce - -lemma contour_integral_reversepath: - assumes "valid_path g" - shows "contour_integral (reversepath g) f = - (contour_integral g f)" -proof (cases "f contour_integrable_on g") - case True then show ?thesis - by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) -next - case False then have "\ f contour_integrable_on (reversepath g)" - by (simp add: assms contour_integrable_reversepath_eq) - with False show ?thesis by (simp add: not_integrable_contour_integral) -qed - - -subsection\<^marker>\tag unimportant\ \Joining two paths together\ - -lemma valid_path_join: - assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" - shows "valid_path(g1 +++ g2)" -proof - - have "g1 1 = g2 0" - using assms by (auto simp: pathfinish_def pathstart_def) - moreover have "(g1 \ (\x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" - apply (rule piecewise_C1_differentiable_compose) - using assms - apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) - apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) - done - moreover have "(g2 \ (\x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" - apply (rule piecewise_C1_differentiable_compose) - using assms unfolding valid_path_def piecewise_C1_differentiable_on_def - by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI - simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) - ultimately show ?thesis - apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) - apply (rule piecewise_C1_differentiable_cases) - apply (auto simp: o_def) - done -qed - -lemma valid_path_join_D1: - fixes g1 :: "real \ 'a::real_normed_field" - shows "valid_path (g1 +++ g2) \ valid_path g1" - unfolding valid_path_def - by (rule piecewise_C1_differentiable_D1) - -lemma valid_path_join_D2: - fixes g2 :: "real \ 'a::real_normed_field" - shows "\valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\ \ valid_path g2" - unfolding valid_path_def - by (rule piecewise_C1_differentiable_D2) - -lemma valid_path_join_eq [simp]: - fixes g2 :: "real \ 'a::real_normed_field" - shows "pathfinish g1 = pathstart g2 \ (valid_path(g1 +++ g2) \ valid_path g1 \ valid_path g2)" - using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast - -lemma has_contour_integral_join: - assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" - "valid_path g1" "valid_path g2" - shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" -proof - - obtain s1 s2 - where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" - and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" - using assms - by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have 1: "((\x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" - and 2: "((\x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" - using assms - by (auto simp: has_contour_integral) - have i1: "((\x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" - and i2: "((\x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" - using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] - has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] - by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) - have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = - 2 *\<^sub>R vector_derivative g1 (at (z*2))" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\z - 1/2\"]]) - apply (simp_all add: dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) - apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - using s1 - apply (auto simp: algebra_simps vector_derivative_works) - done - have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = - 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2 (2*x - 1))" and d = "\z - 1/2\"]]) - apply (simp_all add: dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) - apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - using s2 - apply (auto simp: algebra_simps vector_derivative_works) - done - have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" - apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) ((*)2 -` s1)"]) - using s1 - apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) - apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) - done - moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" - apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) - using s2 - apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) - apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) - done - ultimately - show ?thesis - apply (simp add: has_contour_integral) - apply (rule has_integral_combine [where c = "1/2"], auto) - done -qed - -lemma contour_integrable_joinI: - assumes "f contour_integrable_on g1" "f contour_integrable_on g2" - "valid_path g1" "valid_path g2" - shows "f contour_integrable_on (g1 +++ g2)" - using assms - by (meson has_contour_integral_join contour_integrable_on_def) - -lemma contour_integrable_joinD1: - assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" - shows "f contour_integrable_on g1" -proof - - obtain s1 - where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have "(\x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" - using assms - apply (auto simp: contour_integrable_on) - apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) - apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) - done - then have *: "(\x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" - by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) - have g1: "\0 < z; z < 1; z \ s1\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = - 2 *\<^sub>R vector_derivative g1 (at z)" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\(z-1)/2\"]]) - apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x*2" 2 _ g1, simplified o_def]) - using s1 - apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - done - show ?thesis - using s1 - apply (auto simp: contour_integrable_on) - apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) - apply (auto simp: joinpaths_def scaleR_conv_of_real g1) - done -qed - -lemma contour_integrable_joinD2: - assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" - shows "f contour_integrable_on g2" -proof - - obtain s2 - where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have "(\x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" - using assms - apply (auto simp: contour_integrable_on) - apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) - apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) - apply (simp add: image_affinity_atLeastAtMost_diff) - done - then have *: "(\x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) - integrable_on {0..1}" - by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) - have g2: "\0 < z; z < 1; z \ s2\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = - 2 *\<^sub>R vector_derivative g2 (at z)" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2(2*x-1))" and d = "\z/2\"]]) - apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x*2-1" 2 _ g2, simplified o_def]) - using s2 - apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left - vector_derivative_works add_divide_distrib) - done - show ?thesis - using s2 - apply (auto simp: contour_integrable_on) - apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) - apply (auto simp: joinpaths_def scaleR_conv_of_real g2) - done -qed - -lemma contour_integrable_join [simp]: - shows - "\valid_path g1; valid_path g2\ - \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" -using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast - -lemma contour_integral_join [simp]: - shows - "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ - \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" - by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) - - -subsection\<^marker>\tag unimportant\ \Shifting the starting point of a (closed) path\ - -lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" - by (auto simp: shiftpath_def) - -lemma valid_path_shiftpath [intro]: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "valid_path(shiftpath a g)" - using assms - apply (auto simp: valid_path_def shiftpath_alt_def) - apply (rule piecewise_C1_differentiable_cases) - apply (auto simp: algebra_simps) - apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) - apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) - apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) - apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) - done - -lemma has_contour_integral_shiftpath: - assumes f: "(f has_contour_integral i) g" "valid_path g" - and a: "a \ {0..1}" - shows "(f has_contour_integral i) (shiftpath a g)" -proof - - obtain s - where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have *: "((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" - using assms by (auto simp: has_contour_integral) - then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + - integral {0..a} (\x. f (g x) * vector_derivative g (at x))" - apply (rule has_integral_unique) - apply (subst add.commute) - apply (subst integral_combine) - using assms * integral_unique by auto - { fix x - have "0 \ x \ x + a < 1 \ x \ (\x. x - a) ` s \ - vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" - unfolding shiftpath_def - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x))" and d = "dist(1-a) x"]]) - apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x+a" 1 _ g, simplified o_def scaleR_one]) - apply (intro derivative_eq_intros | simp)+ - using g - apply (drule_tac x="x+a" in bspec) - using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) - done - } note vd1 = this - { fix x - have "1 < x + a \ x \ 1 \ x \ (\x. x - a + 1) ` s \ - vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" - unfolding shiftpath_def - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x-1))" and d = "dist (1-a) x"]]) - apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x+a-1" 1 _ g, simplified o_def scaleR_one]) - apply (intro derivative_eq_intros | simp)+ - using g - apply (drule_tac x="x+a-1" in bspec) - using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) - done - } note vd2 = this - have va1: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" - using * a by (fastforce intro: integrable_subinterval_real) - have v0a: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" - apply (rule integrable_subinterval_real) - using * a by auto - have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) - has_integral integral {a..1} (\x. f (g x) * vector_derivative g (at x))) {0..1 - a}" - apply (rule has_integral_spike_finite - [where S = "{1-a} \ (\x. x-a) ` s" and f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) - using s apply blast - using a apply (auto simp: algebra_simps vd1) - apply (force simp: shiftpath_def add.commute) - using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]] - apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute) - done - moreover - have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) - has_integral integral {0..a} (\x. f (g x) * vector_derivative g (at x))) {1 - a..1}" - apply (rule has_integral_spike_finite - [where S = "{1-a} \ (\x. x-a+1) ` s" and f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) - using s apply blast - using a apply (auto simp: algebra_simps vd2) - apply (force simp: shiftpath_def add.commute) - using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] - apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified]) - apply (simp add: algebra_simps) - done - ultimately show ?thesis - using a - by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) -qed - -lemma has_contour_integral_shiftpath_D: - assumes "(f has_contour_integral i) (shiftpath a g)" - "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "(f has_contour_integral i) g" -proof - - obtain s - where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - { fix x - assume x: "0 < x" "x < 1" "x \ s" - then have gx: "g differentiable at x" - using g by auto - have "vector_derivative g (at x within {0..1}) = - vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" - apply (rule vector_derivative_at_within_ivl - [OF has_vector_derivative_transform_within_open - [where f = "(shiftpath (1 - a) (shiftpath a g))" and S = "{0<..<1}-s"]]) - using s g assms x - apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath - at_within_interior [of _ "{0..1}"] vector_derivative_works [symmetric]) - apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"]) - apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) - done - } note vd = this - have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" - using assms by (auto intro!: has_contour_integral_shiftpath) - show ?thesis - apply (simp add: has_contour_integral_def) - apply (rule has_integral_spike_finite [of "{0,1} \ s", OF _ _ fi [unfolded has_contour_integral_def]]) - using s assms vd - apply (auto simp: Path_Connected.shiftpath_shiftpath) - done -qed - -lemma has_contour_integral_shiftpath_eq: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "(f has_contour_integral i) (shiftpath a g) \ (f has_contour_integral i) g" - using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast - -lemma contour_integrable_on_shiftpath_eq: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" -using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto - -lemma contour_integral_shiftpath: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "contour_integral (shiftpath a g) f = contour_integral g f" - using assms - by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) - - -subsection\<^marker>\tag unimportant\ \More about straight-line paths\ - -lemma has_vector_derivative_linepath_within: - "(linepath a b has_vector_derivative (b - a)) (at x within s)" -apply (simp add: linepath_def has_vector_derivative_def algebra_simps) -apply (rule derivative_eq_intros | simp)+ -done - -lemma vector_derivative_linepath_within: - "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" - apply (rule vector_derivative_within_cbox [of 0 "1::real", simplified]) - apply (auto simp: has_vector_derivative_linepath_within) - done - -lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" - by (simp add: has_vector_derivative_linepath_within vector_derivative_at) - -lemma valid_path_linepath [iff]: "valid_path (linepath a b)" - apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) - apply (rule_tac x="{}" in exI) - apply (simp add: differentiable_on_def differentiable_def) - using has_vector_derivative_def has_vector_derivative_linepath_within - apply (fastforce simp add: continuous_on_eq_continuous_within) - done - -lemma has_contour_integral_linepath: - shows "(f has_contour_integral i) (linepath a b) \ - ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" - by (simp add: has_contour_integral) - -lemma linepath_in_path: - shows "x \ {0..1} \ linepath a b x \ closed_segment a b" - by (auto simp: segment linepath_def) - -lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" - by (auto simp: segment linepath_def) - -lemma linepath_in_convex_hull: - fixes x::real - assumes a: "a \ convex hull s" - and b: "b \ convex hull s" - and x: "0\x" "x\1" - shows "linepath a b x \ convex hull s" - apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) - using x - apply (auto simp: linepath_image_01 [symmetric]) - done - -lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" - by (simp add: linepath_def) - -lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" - by (simp add: linepath_def) - -lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" - by (simp add: has_contour_integral_linepath) - -lemma has_contour_integral_trivial_iff [simp]: "(f has_contour_integral i) (linepath a a) \ i=0" - using has_contour_integral_unique by blast - -lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" - using has_contour_integral_trivial contour_integral_unique by blast - -lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" - by (auto simp: linepath_def) - -lemma bounded_linear_linepath: - assumes "bounded_linear f" - shows "f (linepath a b x) = linepath (f a) (f b) x" -proof - - interpret f: bounded_linear f by fact - show ?thesis by (simp add: linepath_def f.add f.scale) -qed - -lemma bounded_linear_linepath': - assumes "bounded_linear f" - shows "f \ linepath a b = linepath (f a) (f b)" - using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) - -lemma cnj_linepath: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" - by (simp add: linepath_def) - -lemma cnj_linepath': "cnj \ linepath a b = linepath (cnj a) (cnj b)" - by (simp add: linepath_def fun_eq_iff) - -subsection\Relation to subpath construction\ - -lemma valid_path_subpath: - fixes g :: "real \ 'a :: real_normed_vector" - assumes "valid_path g" "u \ {0..1}" "v \ {0..1}" - shows "valid_path(subpath u v g)" -proof (cases "v=u") - case True - then show ?thesis - unfolding valid_path_def subpath_def - by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) -next - case False - have "(g \ (\x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" - apply (rule piecewise_C1_differentiable_compose) - apply (simp add: C1_differentiable_imp_piecewise) - apply (simp add: image_affinity_atLeastAtMost) - using assms False - apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) - apply (subst Int_commute) - apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) - done - then show ?thesis - by (auto simp: o_def valid_path_def subpath_def) -qed - -lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" - by (simp add: has_contour_integral subpath_def) - -lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" - using has_contour_integral_subpath_refl contour_integrable_on_def by blast - -lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" - by (simp add: contour_integral_unique) - -lemma has_contour_integral_subpath: - assumes f: "f contour_integrable_on g" and g: "valid_path g" - and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) - (subpath u v g)" -proof (cases "v=u") - case True - then show ?thesis - using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) -next - case False - obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" - using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast - have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) - has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) - {0..1}" - using f uv - apply (simp add: contour_integrable_on subpath_def has_contour_integral) - apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) - apply (simp_all add: has_integral_integral) - apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) - apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) - apply (simp add: divide_simps False) - done - { fix x - have "x \ {0..1} \ - x \ (\t. (v-u) *\<^sub>R t + u) -` s \ - vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" - apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) - apply (intro derivative_eq_intros | simp)+ - apply (cut_tac s [of "(v - u) * x + u"]) - using uv mult_left_le [of x "v-u"] - apply (auto simp: vector_derivative_works) - done - } note vd = this - show ?thesis - apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) - using fs assms - apply (simp add: False subpath_def has_contour_integral) - apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) - apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) - done -qed - -lemma contour_integrable_subpath: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" - shows "f contour_integrable_on (subpath u v g)" - apply (cases u v rule: linorder_class.le_cases) - apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) - apply (subst reversepath_subpath [symmetric]) - apply (rule contour_integrable_reversepath) - using assms apply (blast intro: valid_path_subpath) - apply (simp add: contour_integrable_on_def) - using assms apply (blast intro: has_contour_integral_subpath) - done - -lemma has_integral_contour_integral_subpath: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "(((\x. f(g x) * vector_derivative g (at x))) - has_integral contour_integral (subpath u v g) f) {u..v}" - using assms - apply (auto simp: has_integral_integrable_integral) - apply (rule integrable_on_subcbox [where a=u and b=v and S = "{0..1}", simplified]) - apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) - done - -lemma contour_integral_subcontour_integral: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "contour_integral (subpath u v g) f = - integral {u..v} (\x. f(g x) * vector_derivative g (at x))" - using assms has_contour_integral_subpath contour_integral_unique by blast - -lemma contour_integral_subpath_combine_less: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" - "u {0..1}" "v \ {0..1}" "w \ {0..1}" - shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = - contour_integral (subpath u w g) f" -proof (cases "u\v \ v\w \ u\w") - case True - have *: "subpath v u g = reversepath(subpath u v g) \ - subpath w u g = reversepath(subpath u w g) \ - subpath w v g = reversepath(subpath v w g)" - by (auto simp: reversepath_subpath) - have "u < v \ v < w \ - u < w \ w < v \ - v < u \ u < w \ - v < w \ w < u \ - w < u \ u < v \ - w < v \ v < u" - using True assms by linarith - with assms show ?thesis - using contour_integral_subpath_combine_less [of f g u v w] - contour_integral_subpath_combine_less [of f g u w v] - contour_integral_subpath_combine_less [of f g v u w] - contour_integral_subpath_combine_less [of f g v w u] - contour_integral_subpath_combine_less [of f g w u v] - contour_integral_subpath_combine_less [of f g w v u] - apply simp - apply (elim disjE) - apply (auto simp: * contour_integral_reversepath contour_integrable_subpath - valid_path_subpath algebra_simps) - done -next - case False - then show ?thesis - apply (auto) - using assms - by (metis eq_neg_iff_add_eq_0 contour_integral_reversepath reversepath_subpath valid_path_subpath) -qed - -lemma contour_integral_integral: - "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" - by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) - -lemma contour_integral_cong: - assumes "g = g'" "\x. x \ path_image g \ f x = f' x" - shows "contour_integral g f = contour_integral g' f'" - unfolding contour_integral_integral using assms - by (intro integral_cong) (auto simp: path_image_def) - - -text \Contour integral along a segment on the real axis\ - -lemma has_contour_integral_linepath_Reals_iff: - fixes a b :: complex and f :: "complex \ complex" - assumes "a \ Reals" "b \ Reals" "Re a < Re b" - shows "(f has_contour_integral I) (linepath a b) \ - ((\x. f (of_real x)) has_integral I) {Re a..Re b}" -proof - - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" - by (simp_all add: complex_eq_iff) - from assms have "a \ b" by auto - have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ - ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" - by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) - (insert assms, simp_all add: field_simps scaleR_conv_of_real) - also have "(\x. f (a + b * of_real x - a * of_real x)) = - (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" - using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) - also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ - ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms - by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) - also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def - by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) - finally show ?thesis by simp -qed - -lemma contour_integrable_linepath_Reals_iff: - fixes a b :: complex and f :: "complex \ complex" - assumes "a \ Reals" "b \ Reals" "Re a < Re b" - shows "(f contour_integrable_on linepath a b) \ - (\x. f (of_real x)) integrable_on {Re a..Re b}" - using has_contour_integral_linepath_Reals_iff[OF assms, of f] - by (auto simp: contour_integrable_on_def integrable_on_def) - -lemma contour_integral_linepath_Reals_eq: - fixes a b :: complex and f :: "complex \ complex" - assumes "a \ Reals" "b \ Reals" "Re a < Re b" - shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" -proof (cases "f contour_integrable_on linepath a b") - case True - thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] - using has_contour_integral_integral has_contour_integral_unique by blast -next - case False - thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] - by (simp add: not_integrable_contour_integral not_integrable_integral) -qed - - - -text\Cauchy's theorem where there's a primitive\ - -lemma contour_integral_primitive_lemma: - fixes f :: "complex \ complex" and g :: "real \ complex" - assumes "a \ b" - and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" - shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) - has_integral (f(g b) - f(g a))) {a..b}" -proof - - obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" - using assms by (auto simp: piecewise_differentiable_on_def) - have cfg: "continuous_on {a..b} (\x. f (g x))" - apply (rule continuous_on_compose [OF cg, unfolded o_def]) - using assms - apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) - done - { fix x::real - assume a: "a < x" and b: "x < b" and xk: "x \ k" - then have "g differentiable at x within {a..b}" - using k by (simp add: differentiable_at_withinI) - then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" - by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) - then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" - by (simp add: has_vector_derivative_def scaleR_conv_of_real) - have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" - using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) - then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" - by (simp add: has_field_derivative_def) - have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" - using diff_chain_within [OF gdiff fdiff] - by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) - } note * = this - show ?thesis - apply (rule fundamental_theorem_of_calculus_interior_strong) - using k assms cfg * - apply (auto simp: at_within_Icc_at) - done -qed - -lemma contour_integral_primitive: - assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "valid_path g" "path_image g \ s" - shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" - using assms - apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) - apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) - done - -corollary Cauchy_theorem_primitive: - assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "valid_path g" "path_image g \ s" "pathfinish g = pathstart g" - shows "(f' has_contour_integral 0) g" - using assms - by (metis diff_self contour_integral_primitive) - -text\Existence of path integral for continuous function\ -lemma contour_integrable_continuous_linepath: - assumes "continuous_on (closed_segment a b) f" - shows "f contour_integrable_on (linepath a b)" -proof - - have "continuous_on {0..1} ((\x. f x * (b - a)) \ linepath a b)" - apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) - apply (rule continuous_intros | simp add: assms)+ - done - then show ?thesis - apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) - apply (rule integrable_continuous [of 0 "1::real", simplified]) - apply (rule continuous_on_eq [where f = "\x. f(linepath a b x)*(b - a)"]) - apply (auto simp: vector_derivative_linepath_within) - done -qed - -lemma has_field_der_id: "((\x. x\<^sup>2 / 2) has_field_derivative x) (at x)" - by (rule has_derivative_imp_has_field_derivative) - (rule derivative_intros | simp)+ - -lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" - apply (rule contour_integral_unique) - using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] - apply (auto simp: field_simps has_field_der_id) - done - -lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" - by (simp add: contour_integrable_continuous_linepath) - -lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" - by (simp add: contour_integrable_continuous_linepath) - -subsection\<^marker>\tag unimportant\ \Arithmetical combining theorems\ - -lemma has_contour_integral_neg: - "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" - by (simp add: has_integral_neg has_contour_integral_def) - -lemma has_contour_integral_add: - "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ - \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" - by (simp add: has_integral_add has_contour_integral_def algebra_simps) - -lemma has_contour_integral_diff: - "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ - \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" - by (simp add: has_integral_diff has_contour_integral_def algebra_simps) - -lemma has_contour_integral_lmul: - "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" -apply (simp add: has_contour_integral_def) -apply (drule has_integral_mult_right) -apply (simp add: algebra_simps) -done - -lemma has_contour_integral_rmul: - "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" -apply (drule has_contour_integral_lmul) -apply (simp add: mult.commute) -done - -lemma has_contour_integral_div: - "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" - by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) - -lemma has_contour_integral_eq: - "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" -apply (simp add: path_image_def has_contour_integral_def) -by (metis (no_types, lifting) image_eqI has_integral_eq) - -lemma has_contour_integral_bound_linepath: - assumes "(f has_contour_integral i) (linepath a b)" - "0 \ B" "\x. x \ closed_segment a b \ norm(f x) \ B" - shows "norm i \ B * norm(b - a)" -proof - - { fix x::real - assume x: "0 \ x" "x \ 1" - have "norm (f (linepath a b x)) * - norm (vector_derivative (linepath a b) (at x within {0..1})) \ B * norm (b - a)" - by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) - } note * = this - have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" - apply (rule has_integral_bound - [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) - using assms * unfolding has_contour_integral_def - apply (auto simp: norm_mult) - done - then show ?thesis - by (auto simp: content_real) -qed - -(*UNUSED -lemma has_contour_integral_bound_linepath_strong: - fixes a :: real and f :: "complex \ real" - assumes "(f has_contour_integral i) (linepath a b)" - "finite k" - "0 \ B" "\x::real. x \ closed_segment a b - k \ norm(f x) \ B" - shows "norm i \ B*norm(b - a)" -*) - -lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" - unfolding has_contour_integral_linepath - by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) - -lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" - by (simp add: has_contour_integral_def) - -lemma has_contour_integral_is_0: - "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" - by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto - -lemma has_contour_integral_sum: - "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ - \ ((\x. sum (\a. f a x) s) has_contour_integral sum i s) p" - by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) - -subsection\<^marker>\tag unimportant\ \Operations on path integrals\ - -lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" - by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) - -lemma contour_integral_neg: - "f contour_integrable_on g \ contour_integral g (\x. -(f x)) = -(contour_integral g f)" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) - -lemma contour_integral_add: - "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = - contour_integral g f1 + contour_integral g f2" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) - -lemma contour_integral_diff: - "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = - contour_integral g f1 - contour_integral g f2" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) - -lemma contour_integral_lmul: - shows "f contour_integrable_on g - \ contour_integral g (\x. c * f x) = c*contour_integral g f" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) - -lemma contour_integral_rmul: - shows "f contour_integrable_on g - \ contour_integral g (\x. f x * c) = contour_integral g f * c" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) - -lemma contour_integral_div: - shows "f contour_integrable_on g - \ contour_integral g (\x. f x / c) = contour_integral g f / c" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) - -lemma contour_integral_eq: - "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" - apply (simp add: contour_integral_def) - using has_contour_integral_eq - by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) - -lemma contour_integral_eq_0: - "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" - by (simp add: has_contour_integral_is_0 contour_integral_unique) - -lemma contour_integral_bound_linepath: - shows - "\f contour_integrable_on (linepath a b); - 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ - \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" - apply (rule has_contour_integral_bound_linepath [of f]) - apply (auto simp: has_contour_integral_integral) - done - -lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" - by (simp add: contour_integral_unique has_contour_integral_0) - -lemma contour_integral_sum: - "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ - \ contour_integral p (\x. sum (\a. f a x) s) = sum (\a. contour_integral p (f a)) s" - by (auto simp: contour_integral_unique has_contour_integral_sum has_contour_integral_integral) - -lemma contour_integrable_eq: - "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" - unfolding contour_integrable_on_def - by (metis has_contour_integral_eq) - - -subsection\<^marker>\tag unimportant\ \Arithmetic theorems for path integrability\ - -lemma contour_integrable_neg: - "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" - using has_contour_integral_neg contour_integrable_on_def by blast - -lemma contour_integrable_add: - "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" - using has_contour_integral_add contour_integrable_on_def - by fastforce - -lemma contour_integrable_diff: - "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" - using has_contour_integral_diff contour_integrable_on_def - by fastforce - -lemma contour_integrable_lmul: - "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" - using has_contour_integral_lmul contour_integrable_on_def - by fastforce - -lemma contour_integrable_rmul: - "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" - using has_contour_integral_rmul contour_integrable_on_def - by fastforce - -lemma contour_integrable_div: - "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" - using has_contour_integral_div contour_integrable_on_def - by fastforce - -lemma contour_integrable_sum: - "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ - \ (\x. sum (\a. f a x) s) contour_integrable_on p" - unfolding contour_integrable_on_def - by (metis has_contour_integral_sum) - - -subsection\<^marker>\tag unimportant\ \Reversing a path integral\ - -lemma has_contour_integral_reverse_linepath: - "(f has_contour_integral i) (linepath a b) - \ (f has_contour_integral (-i)) (linepath b a)" - using has_contour_integral_reversepath valid_path_linepath by fastforce - -lemma contour_integral_reverse_linepath: - "continuous_on (closed_segment a b) f - \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" -apply (rule contour_integral_unique) -apply (rule has_contour_integral_reverse_linepath) -by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) - - -(* Splitting a path integral in a flat way.*) - -lemma has_contour_integral_split: - assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "(f has_contour_integral (i + j)) (linepath a b)" -proof (cases "k = 0 \ k = 1") - case True - then show ?thesis - using assms by auto -next - case False - then have k: "0 < k" "k < 1" "complex_of_real k \ 1" - using assms by auto - have c': "c = k *\<^sub>R (b - a) + a" - by (metis diff_add_cancel c) - have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" - by (simp add: algebra_simps c') - { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" - have **: "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" - using False apply (simp add: c' algebra_simps) - apply (simp add: real_vector.scale_left_distrib [symmetric] field_split_simps) - done - have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" - using k has_integral_affinity01 [OF *, of "inverse k" "0"] - apply (simp add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) - apply (auto dest: has_integral_cmul [where c = "inverse k"]) - done - } note fi = this - { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" - have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" - using k - apply (simp add: c' field_simps) - apply (simp add: scaleR_conv_of_real divide_simps) - apply (simp add: field_simps) - done - have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" - using k has_integral_affinity01 [OF *, of "inverse(1 - k)" "-(k/(1 - k))"] - apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) - apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) - done - } note fj = this - show ?thesis - using f k - apply (simp add: has_contour_integral_linepath) - apply (simp add: linepath_def) - apply (rule has_integral_combine [OF _ _ fi fj], simp_all) - done -qed - -lemma continuous_on_closed_segment_transform: - assumes f: "continuous_on (closed_segment a b) f" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "continuous_on (closed_segment a c) f" -proof - - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" - using c by (simp add: algebra_simps) - have "closed_segment a c \ closed_segment a b" - by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) - then show "continuous_on (closed_segment a c) f" - by (rule continuous_on_subset [OF f]) -qed - -lemma contour_integral_split: - assumes f: "continuous_on (closed_segment a b) f" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" -proof - - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" - using c by (simp add: algebra_simps) - have "closed_segment a c \ closed_segment a b" - by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) - moreover have "closed_segment c b \ closed_segment a b" - by (metis c' ends_in_segment(2) in_segment(1) k subset_closed_segment) - ultimately - have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" - by (auto intro: continuous_on_subset [OF f]) - show ?thesis - by (rule contour_integral_unique) (meson "*" c contour_integrable_continuous_linepath has_contour_integral_integral has_contour_integral_split k) -qed - -lemma contour_integral_split_linepath: - assumes f: "continuous_on (closed_segment a b) f" - and c: "c \ closed_segment a b" - shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" - using c by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) - text\The special case of midpoints used in the main quadrisection\ lemma has_contour_integral_midpoint: assumes "(f has_contour_integral i) (linepath a (midpoint a b))" "(f has_contour_integral j) (linepath (midpoint a b) b)" shows "(f has_contour_integral (i + j)) (linepath a b)" apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"]) using assms apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) done lemma contour_integral_midpoint: "continuous_on (closed_segment a b) f \ contour_integral (linepath a b) f = contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f" apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"]) apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) done text\A couple of special case lemmas that are useful below\ lemma triangle_linear_has_chain_integral: "((\x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" apply (rule Cauchy_theorem_primitive [of UNIV "\x. m/2 * x^2 + d*x"]) apply (auto intro!: derivative_eq_intros) done lemma has_chain_integral_chain_integral3: "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d) \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i" apply (subst contour_integral_unique [symmetric], assumption) apply (drule has_contour_integral_integrable) apply (simp add: valid_path_join) done lemma has_chain_integral_chain_integral4: "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e) \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i" apply (subst contour_integral_unique [symmetric], assumption) apply (drule has_contour_integral_integrable) apply (simp add: valid_path_join) done subsection\Reversing the order in a double path integral\ text\The condition is stronger than needed but it's often true in typical situations\ lemma fst_im_cbox [simp]: "cbox c d \ {} \ (fst ` cbox (a,c) (b,d)) = cbox a b" by (auto simp: cbox_Pair_eq) lemma snd_im_cbox [simp]: "cbox a b \ {} \ (snd ` cbox (a,c) (b,d)) = cbox c d" by (auto simp: cbox_Pair_eq) proposition contour_integral_swap: assumes fcon: "continuous_on (path_image g \ path_image h) (\(y1,y2). f y1 y2)" and vp: "valid_path g" "valid_path h" and gvcon: "continuous_on {0..1} (\t. vector_derivative g (at t))" and hvcon: "continuous_on {0..1} (\t. vector_derivative h (at t))" shows "contour_integral g (\w. contour_integral h (f w)) = contour_integral h (\z. contour_integral g (\w. f w z))" proof - have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have fgh1: "\x. (\t. f (g x) (h t)) = (\(y1,y2). f y1 y2) \ (\t. (g x, h t))" by (rule ext) simp have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) \ (\t. (g t, h x))" by (rule ext) simp have fcon_im1: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g x, h t)) ` {0..1}) (\(x, y). f x y)" by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) have fcon_im2: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g t, h x)) ` {0..1}) (\(x, y). f x y)" by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) have "\y. y \ {0..1} \ continuous_on {0..1} (\x. f (g x) (h y))" by (subst fgh2) (rule fcon_im2 gcon continuous_intros | simp)+ then have vdg: "\y. y \ {0..1} \ (\x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}" using continuous_on_mult gvcon integrable_continuous_real by blast have "(\z. vector_derivative g (at (fst z))) = (\x. vector_derivative g (at x)) \ fst" by auto then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\x. vector_derivative g (at (fst x)))" apply (rule ssubst) apply (rule continuous_intros | simp add: gvcon)+ done have "(\z. vector_derivative h (at (snd z))) = (\x. vector_derivative h (at x)) \ snd" by auto then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" apply (rule ssubst) apply (rule continuous_intros | simp add: hvcon)+ done have "(\x. f (g (fst x)) (h (snd x))) = (\(y1,y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w))" by auto then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" apply (rule ssubst) apply (rule gcon hcon continuous_intros | simp)+ apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) done have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = integral {0..1} (\x. contour_integral h (\y. f (g x) y * vector_derivative g (at x)))" proof (rule integral_cong [OF contour_integral_rmul [symmetric]]) show "\x. x \ {0..1} \ f (g x) contour_integrable_on h" unfolding contour_integrable_on apply (rule integrable_continuous_real) apply (rule continuous_on_mult [OF _ hvcon]) apply (subst fgh1) apply (rule fcon_im1 hcon continuous_intros | simp)+ done qed also have "\ = integral {0..1} (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" unfolding contour_integral_integral apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ unfolding integral_mult_left [symmetric] apply (simp only: mult_ac) done also have "\ = contour_integral h (\z. contour_integral g (\w. f w z))" unfolding contour_integral_integral apply (rule integral_cong) unfolding integral_mult_left [symmetric] apply (simp add: algebra_simps) done finally show ?thesis by (simp add: contour_integral_integral) qed subsection\<^marker>\tag unimportant\ \The key quadrisection step\ lemma norm_sum_half: assumes "norm(a + b) \ e" shows "norm a \ e/2 \ norm b \ e/2" proof - have "e \ norm (- a - b)" by (simp add: add.commute assms norm_minus_commute) thus ?thesis using norm_triangle_ineq4 order_trans by fastforce qed lemma norm_sum_lemma: assumes "e \ norm (a + b + c + d)" shows "e / 4 \ norm a \ e / 4 \ norm b \ e / 4 \ norm c \ e / 4 \ norm d" proof - have "e \ norm ((a + b) + (c + d))" using assms by (simp add: algebra_simps) then show ?thesis by (auto dest!: norm_sum_half) qed lemma Cauchy_theorem_quadrisection: assumes f: "continuous_on (convex hull {a,b,c}) f" and dist: "dist a b \ K" "dist b c \ K" "dist c a \ K" and e: "e * K^2 \ norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)" shows "\a' b' c'. a' \ convex hull {a,b,c} \ b' \ convex hull {a,b,c} \ c' \ convex hull {a,b,c} \ dist a' b' \ K/2 \ dist b' c' \ K/2 \ dist c' a' \ K/2 \ e * (K/2)^2 \ norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" (is "\x y z. ?\ x y z") proof - note divide_le_eq_numeral1 [simp del] define a' where "a' = midpoint b c" define b' where "b' = midpoint c a" define c' where "c' = midpoint a b" have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ have fcont': "continuous_on (closed_segment c' b') f" "continuous_on (closed_segment a' c') f" "continuous_on (closed_segment b' a') f" unfolding a'_def b'_def c'_def by (rule continuous_on_subset [OF f], metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+ let ?pathint = "\x y. contour_integral(linepath x y) f" have *: "?pathint a b + ?pathint b c + ?pathint c a = (?pathint a c' + ?pathint c' b' + ?pathint b' a) + (?pathint a' c' + ?pathint c' b + ?pathint b a') + (?pathint a' c + ?pathint c b' + ?pathint b' a') + (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" by (simp add: fcont' contour_integral_reverse_linepath) (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc) have [simp]: "\x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2" by (metis left_diff_distrib mult.commute norm_mult_numeral1) have [simp]: "\x y. cmod (x - y) = cmod (y - x)" by (simp add: norm_minus_commute) consider "e * K\<^sup>2 / 4 \ cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" using assms unfolding * by (blast intro: that dest!: norm_sum_lemma) then show ?thesis proof cases case 1 then have "?\ a c' b'" using assms apply (clarsimp simp: c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 2 then have "?\ a' c' b" using assms apply (clarsimp simp: a'_def c'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 3 then have "?\ a' c b'" using assms apply (clarsimp simp: a'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 4 then have "?\ a' b' c'" using assms apply (clarsimp simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast qed qed subsection\<^marker>\tag unimportant\ \Cauchy's theorem for triangles\ lemma triangle_points_closer: fixes a::complex shows "\x \ convex hull {a,b,c}; y \ convex hull {a,b,c}\ \ norm(x - y) \ norm(a - b) \ norm(x - y) \ norm(b - c) \ norm(x - y) \ norm(c - a)" using simplex_extremal_le [of "{a,b,c}"] by (auto simp: norm_minus_commute) lemma holomorphic_point_small_triangle: assumes x: "x \ S" and f: "continuous_on S f" and cd: "f field_differentiable (at x within S)" and e: "0 < e" shows "\k>0. \a b c. dist a b \ k \ dist b c \ k \ dist c a \ k \ x \ convex hull {a,b,c} \ convex hull {a,b,c} \ S \ norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f) \ e*(dist a b + dist b c + dist c a)^2" (is "\k>0. \a b c. _ \ ?normle a b c") proof - have le_of_3: "\a x y z. \0 \ x*y; 0 \ x*z; 0 \ y*z; a \ (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\ \ a \ e*(x + y + z)^2" by (simp add: algebra_simps power2_eq_square) have disj_le: "\x \ a \ x \ b \ x \ c; 0 \ a; 0 \ b; 0 \ c\ \ x \ a + b + c" for x::real and a b c by linarith have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a" if "convex hull {a, b, c} \ S" for a b c using segments_subset_convex_hull that by (metis continuous_on_subset f contour_integrable_continuous_linepath)+ note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral] { fix f' a b c d assume d: "0 < d" and f': "\y. \cmod (y - x) \ d; y \ S\ \ cmod (f y - f x - f' * (y - x)) \ e * cmod (y - x)" and le: "cmod (a - b) \ d" "cmod (b - c) \ d" "cmod (c - a) \ d" and xc: "x \ convex hull {a, b, c}" and S: "convex hull {a, b, c} \ S" have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = contour_integral (linepath a b) (\y. f y - f x - f'*(y - x)) + contour_integral (linepath b c) (\y. f y - f x - f'*(y - x)) + contour_integral (linepath c a) (\y. f y - f x - f'*(y - x))" apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF S]) apply (simp add: field_simps) done { fix y assume yc: "y \ convex hull {a,b,c}" have "cmod (f y - f x - f' * (y - x)) \ e*norm(y - x)" proof (rule f') show "cmod (y - x) \ d" by (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans) qed (use S yc in blast) also have "\ \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" by (simp add: yc e xc disj_le [OF triangle_points_closer]) finally have "cmod (f y - f x - f' * (y - x)) \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" . } note cm_le = this have "?normle a b c" unfolding dist_norm pa apply (rule le_of_3) using f' xc S e apply simp_all apply (intro norm_triangle_le add_mono path_bound) apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc) apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+ done } note * = this show ?thesis using cd e apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def) apply (clarify dest!: spec mp) using * unfolding dist_norm apply blast done qed text\Hence the most basic theorem for a triangle.\ locale Chain = fixes x0 At Follows assumes At0: "At x0 0" and AtSuc: "\x n. At x n \ \x'. At x' (Suc n) \ Follows x' x" begin primrec f where "f 0 = x0" | "f (Suc n) = (SOME x. At x (Suc n) \ Follows x (f n))" lemma At: "At (f n) n" proof (induct n) case 0 show ?case by (simp add: At0) next case (Suc n) show ?case by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex) qed lemma Follows: "Follows (f(Suc n)) (f n)" by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex) declare f.simps(2) [simp del] end lemma Chain3: assumes At0: "At x0 y0 z0 0" and AtSuc: "\x y z n. At x y z n \ \x' y' z'. At x' y' z' (Suc n) \ Follows x' y' z' x y z" obtains f g h where "f 0 = x0" "g 0 = y0" "h 0 = z0" "\n. At (f n) (g n) (h n) n" "\n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)" proof - interpret three: Chain "(x0,y0,z0)" "\(x,y,z). At x y z" "\(x',y',z'). \(x,y,z). Follows x' y' z' x y z" apply unfold_locales using At0 AtSuc by auto show ?thesis apply (rule that [of "\n. fst (three.f n)" "\n. fst (snd (three.f n))" "\n. snd (snd (three.f n))"]) using three.At three.Follows apply simp_all apply (simp_all add: split_beta') done qed proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle: assumes "f holomorphic_on (convex hull {a,b,c})" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof - have contf: "continuous_on (convex hull {a,b,c}) f" by (metis assms holomorphic_on_imp_continuous_on) let ?pathint = "\x y. contour_integral(linepath x y) f" { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" define e where "e = norm y / K^2" have K1: "K \ 1" by (simp add: K_def max.coboundedI1) then have K: "K > 0" by linarith have [iff]: "dist a b \ K" "dist b c \ K" "dist c a \ K" by (simp_all add: K_def) have e: "e > 0" unfolding e_def using ynz K1 by simp define At where "At x y z n \ convex hull {x,y,z} \ convex hull {a,b,c} \ dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" for x y z n have At0: "At a b c 0" using fy by (simp add: At_def e_def has_chain_integral_chain_integral3) { fix x y z n assume At: "At x y z n" then have contf': "continuous_on (convex hull {x,y,z}) f" using contf At_def continuous_on_subset by metis have "\x' y' z'. At x' y' z' (Suc n) \ convex hull {x',y',z'} \ convex hull {x,y,z}" using At Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e] apply (simp add: At_def algebra_simps) apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE) done } note AtSuc = this obtain fa fb fc where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c" and cosb: "\n. convex hull {fa n, fb n, fc n} \ convex hull {a,b,c}" and dist: "\n. dist (fa n) (fb n) \ K/2^n" "\n. dist (fb n) (fc n) \ K/2^n" "\n. dist (fc n) (fa n) \ K/2^n" and no: "\n. norm(?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) \ e * (K/2^n)^2" and conv_le: "\n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \ convex hull {fa n, fb n, fc n}" apply (rule Chain3 [of At, OF At0 AtSuc]) apply (auto simp: At_def) done obtain x where x: "\n. x \ convex hull {fa n, fb n, fc n}" proof (rule bounded_closed_nest) show "\n. closed (convex hull {fa n, fb n, fc n})" by (simp add: compact_imp_closed finite_imp_compact_convex_hull) show "\m n. m \ n \ convex hull {fa n, fb n, fc n} \ convex hull {fa m, fb m, fc m}" by (erule transitive_stepwise_le) (auto simp: conv_le) qed (fastforce intro: finite_imp_bounded_convex_hull)+ then have xin: "x \ convex hull {a,b,c}" using assms f0 by blast then have fx: "f field_differentiable at x within (convex hull {a,b,c})" using assms holomorphic_on_def by blast { fix k n assume k: "0 < k" and le: "\x' y' z'. \dist x' y' \ k; dist y' z' \ k; dist z' x' \ k; x \ convex hull {x',y',z'}; convex hull {x',y',z'} \ convex hull {a,b,c}\ \ cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10 \ e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2" and Kk: "K / k < 2 ^ n" have "K / 2 ^ n < k" using Kk k by (auto simp: field_simps) then have DD: "dist (fa n) (fb n) \ k" "dist (fb n) (fc n) \ k" "dist (fc n) (fa n) \ k" using dist [of n] k by linarith+ have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ (3 * K / 2 ^ n)\<^sup>2" using dist [of n] e K by (simp add: abs_le_square_iff [symmetric]) have less10: "\x y::real. 0 < x \ y \ 9*x \ y < x*10" by linarith have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ e * (3 * K / 2 ^ n)\<^sup>2" using ynz dle e mult_le_cancel_left_pos by blast also have "\ < cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10" using no [of n] e K apply (simp add: e_def field_simps) apply (simp only: zero_less_norm_iff [symmetric]) done finally have False using le [OF DD x cosb] by auto } then have ?thesis using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e apply clarsimp apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]], force+) done } moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1) segments_subset_convex_hull(3) segments_subset_convex_hull(5)) ultimately show ?thesis using has_contour_integral_integral by fastforce qed subsection\<^marker>\tag unimportant\ \Version needing function holomorphic in interior only\ lemma Cauchy_theorem_flat_lemma: assumes f: "continuous_on (convex hull {a,b,c}) f" and c: "c - a = k *\<^sub>R (b - a)" and k: "0 \ k" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ show ?thesis proof (cases "k \ 1") case True show ?thesis by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc) next case False then show ?thesis using fabc c apply (subst contour_integral_split [of a c f "1/k" b, symmetric]) apply (metis closed_segment_commute fabc(3)) apply (auto simp: k contour_integral_reverse_linepath) done qed qed lemma Cauchy_theorem_flat: assumes f: "continuous_on (convex hull {a,b,c}) f" and c: "c - a = k *\<^sub>R (b - a)" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "0 \ k") case True with assms show ?thesis by (blast intro: Cauchy_theorem_flat_lemma) next case False have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f + contour_integral (linepath c b) f = 0" apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"]) using False c apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps) done ultimately show ?thesis apply (auto simp: contour_integral_reverse_linepath) using add_eq_0_iff by force qed lemma Cauchy_theorem_triangle_interior: assumes contf: "continuous_on (convex hull {a,b,c}) f" and holf: "f holomorphic_on interior (convex hull {a,b,c})" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using contf continuous_on_subset segments_subset_convex_hull by metis+ have "bounded (f ` (convex hull {a,b,c}))" by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf) then obtain B where "0 < B" and Bnf: "\x. x \ convex hull {a,b,c} \ norm (f x) \ B" by (auto simp: dest!: bounded_pos [THEN iffD1]) have "bounded (convex hull {a,b,c})" by (simp add: bounded_convex_hull) then obtain C where C: "0 < C" and Cno: "\y. y \ convex hull {a,b,c} \ norm y < C" using bounded_pos_less by blast then have diff_2C: "norm(x - y) \ 2*C" if x: "x \ convex hull {a, b, c}" and y: "y \ convex hull {a, b, c}" for x y proof - have "cmod x \ C" using x by (meson Cno not_le not_less_iff_gr_or_eq) hence "cmod (x - y) \ C + C" using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans) thus "cmod (x - y) \ 2 * C" by (metis mult_2) qed have contf': "continuous_on (convex hull {b,a,c}) f" using contf by (simp add: insert_commute) { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y" by (rule has_chain_integral_chain_integral3 [OF fy]) have ?thesis proof (cases "c=a \ a=b \ b=c") case True then show ?thesis using Cauchy_theorem_flat [OF contf, of 0] using has_chain_integral_chain_integral3 [OF fy] ynz by (force simp: fabc contour_integral_reverse_linepath) next case False then have car3: "card {a, b, c} = Suc (DIM(complex))" by auto { assume "interior(convex hull {a,b,c}) = {}" then have "collinear{a,b,c}" using interior_convex_hull_eq_empty [OF car3] by (simp add: collinear_3_eq_affine_dependent) with False obtain d where "c \ a" "a \ b" "b \ c" "c - b = d *\<^sub>R (a - b)" by (auto simp: collinear_3 collinear_lemma) then have "False" using False Cauchy_theorem_flat [OF contf'] pi_eq_y ynz by (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath) } then obtain d where d: "d \ interior (convex hull {a, b, c})" by blast { fix d1 assume d1_pos: "0 < d1" and d1: "\x x'. \x\convex hull {a, b, c}; x'\convex hull {a, b, c}; cmod (x' - x) < d1\ \ cmod (f x' - f x) < cmod y / (24 * C)" define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x let ?pathint = "\x y. contour_integral(linepath x y) f" have e: "0 < e" "e \ 1" "e \ d1 / (4 * C)" "e \ cmod y / 24 / C / B" using d1_pos \C>0\ \B>0\ ynz by (simp_all add: e_def) then have eCB: "24 * e * C * B \ cmod y" using \C>0\ \B>0\ by (simp add: field_simps) have e_le_d1: "e * (4 * C) \ d1" using e \C>0\ by (simp add: field_simps) have "shrink a \ interior(convex hull {a,b,c})" "shrink b \ interior(convex hull {a,b,c})" "shrink c \ interior(convex hull {a,b,c})" using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) then have fhp0: "(f has_contour_integral 0) (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))" by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal) then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0" by (simp add: has_chain_integral_chain_integral3) have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)" "f contour_integrable_on linepath (shrink b) (shrink c)" "f contour_integrable_on linepath (shrink c) (shrink a)" using fhp0 by (auto simp: valid_path_join dest: has_contour_integral_integrable) have cmod_shr: "\x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)" using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric]) have sh_eq: "\a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)" by (simp add: algebra_simps) have "cmod y / (24 * C) \ cmod y / cmod (b - a) / 12" using False \C>0\ diff_2C [of b a] ynz by (auto simp: field_split_simps hull_inc) have less_C: "\u \ convex hull {a, b, c}; 0 \ x; x \ 1\ \ x * cmod u < C" for x u apply (cases "x=0", simp add: \0) using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast { fix u v assume uv: "u \ convex hull {a, b, c}" "v \ convex hull {a, b, c}" "u\v" and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)" have shr_uv: "shrink u \ interior(convex hull {a,b,c})" "shrink v \ interior(convex hull {a,b,c})" using d e uv by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) have cmod_fuv: "\x. 0\x \ x\1 \ cmod (f (linepath (shrink u) (shrink v) x)) \ B" using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD]) have By_uv: "B * (12 * (e * cmod (u - v))) \ cmod y" apply (rule order_trans [OF _ eCB]) using e \B>0\ diff_2C [of u v] uv by (auto simp: field_simps) { fix x::real assume x: "0\x" "x\1" have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)" apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0]) using uv x d interior_subset apply (auto simp: hull_inc intro!: less_C) done have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))" by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real) have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1" apply (simp only: ll norm_mult scaleR_diff_right) using \e>0\ cmod_less_4C apply (force intro: norm_triangle_lt less_le_trans [OF _ e_le_d1]) done have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)" using x uv shr_uv cmod_less_dt by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull) also have "\ \ cmod y / cmod (v - u) / 12" using False uv \C>0\ diff_2C [of v u] ynz by (auto simp: field_split_simps hull_inc) finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / cmod (v - u) / 12" by simp then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \ cmod y" using uv False by (auto simp: field_simps) have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ B * (cmod y / 24 / C / B * 2 * C) + 2 * C * (cmod y / 24 / C)" apply (rule add_mono [OF mult_mono]) using By_uv e \0 < B\ \0 < C\ x apply (simp_all add: cmod_fuv cmod_shr cmod_12_le) apply (simp add: field_simps) done also have "\ \ cmod y / 6" by simp finally have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / 6" . } note cmod_diff_le = this have f_uv: "continuous_on (closed_segment u v) f" by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull]) have **: "\f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)" by (simp add: algebra_simps) have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ (B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)) * content (cbox 0 (1::real))" apply (rule has_integral_bound [of _ "\x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)" _ 0 1]) using ynz \0 < B\ \0 < C\ apply (simp_all del: le_divide_eq_numeral1) apply (simp add: has_integral_diff has_contour_integral_linepath [symmetric] has_contour_integral_integral fpi_uv f_uv contour_integrable_continuous_linepath) apply (auto simp: ** norm_triangle_le norm_mult cmod_diff_le simp del: le_divide_eq_numeral1) done also have "\ \ norm y / 6" by simp finally have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ norm y / 6" . } note * = this have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) moreover have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) moreover have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) ultimately have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) + (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a)) \ norm y / 6 + norm y / 6 + norm y / 6" by (metis norm_triangle_le add_mono) also have "\ = norm y / 2" by simp finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) - (?pathint a b + ?pathint b c + ?pathint c a)) \ norm y / 2" by (simp add: algebra_simps) then have "norm(?pathint a b + ?pathint b c + ?pathint c a) \ norm y / 2" by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff) then have "False" using pi_eq_y ynz by auto } moreover have "uniformly_continuous_on (convex hull {a,b,c}) f" by (simp add: contf compact_convex_hull compact_uniformly_continuous) ultimately have "False" unfolding uniformly_continuous_on_def by (force simp: ynz \0 < C\ dist_norm) then show ?thesis .. qed } moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" using fabc contour_integrable_continuous_linepath by auto ultimately show ?thesis using has_contour_integral_integral by fastforce qed subsection\<^marker>\tag unimportant\ \Version allowing finite number of exceptional points\ proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle_cofinite: assumes "continuous_on (convex hull {a,b,c}) f" and "finite S" and "(\x. x \ interior(convex hull {a,b,c}) - S \ f field_differentiable (at x))" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" using assms proof (induction "card S" arbitrary: a b c S rule: less_induct) case (less S a b c) show ?case proof (cases "S={}") case True with less show ?thesis by (fastforce simp: holomorphic_on_def field_differentiable_at_within Cauchy_theorem_triangle_interior) next case False then obtain d S' where d: "S = insert d S'" "d \ S'" by (meson Set.set_insert all_not_in_conv) then show ?thesis proof (cases "d \ convex hull {a,b,c}") case False show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof (rule less.hyps) show "\x. x \ interior (convex hull {a, b, c}) - S' \ f field_differentiable at x" using False d interior_subset by (auto intro!: less.prems) qed (use d less.prems in auto) next case True have *: "convex hull {a, b, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)" proof (rule less.hyps) show "\x. x \ interior (convex hull {a, b, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have *: "convex hull {b, c, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)" proof (rule less.hyps) show "\x. x \ interior (convex hull {b, c, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have *: "convex hull {c, a, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)" proof (rule less.hyps) show "\x. x \ interior (convex hull {c, a, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have "f contour_integrable_on linepath a b" using less.prems abd contour_integrable_joinD1 contour_integrable_on_def by blast moreover have "f contour_integrable_on linepath b c" using less.prems bcd contour_integrable_joinD1 contour_integrable_on_def by blast moreover have "f contour_integrable_on linepath c a" using less.prems cad contour_integrable_joinD1 contour_integrable_on_def by blast ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" by auto { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" have cont_ad: "continuous_on (closed_segment a d) f" by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3)) have cont_bd: "continuous_on (closed_segment b d) f" by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1)) have cont_cd: "continuous_on (closed_segment c d) f" by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2)) have "contour_integral (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))" "contour_integral (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))" "contour_integral (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)" using has_chain_integral_chain_integral3 [OF abd] has_chain_integral_chain_integral3 [OF bcd] has_chain_integral_chain_integral3 [OF cad] by (simp_all add: algebra_simps add_eq_0_iff) then have ?thesis using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce } then show ?thesis using fpi contour_integrable_on_def by blast qed qed qed subsection\<^marker>\tag unimportant\ \Cauchy's theorem for an open starlike set\ lemma starlike_convex_subset: assumes S: "a \ S" "closed_segment b c \ S" and subs: "\x. x \ S \ closed_segment a x \ S" shows "convex hull {a,b,c} \ S" using S apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull) apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE) done lemma triangle_contour_integrals_starlike_primitive: assumes contf: "continuous_on S f" and S: "a \ S" "open S" and x: "x \ S" and subs: "\y. y \ S \ closed_segment a y \ S" and zer: "\b c. closed_segment b c \ S \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)" proof - let ?pathint = "\x y. contour_integral(linepath x y) f" { fix e y assume e: "0 < e" and bxe: "ball x e \ S" and close: "cmod (y - x) < e" have y: "y \ S" using bxe close by (force simp: dist_norm norm_minus_commute) have cont_ayf: "continuous_on (closed_segment a y) f" using contf continuous_on_subset subs y by blast have xys: "closed_segment x y \ S" apply (rule order_trans [OF _ bxe]) using close by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound) have "?pathint a y - ?pathint a x = ?pathint x y" using zer [OF xys] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force } note [simp] = this { fix e::real assume e: "0 < e" have cont_atx: "continuous (at x) f" using x S contf continuous_on_eq_continuous_at by blast then obtain d1 where d1: "d1>0" and d1_less: "\y. cmod (y - x) < d1 \ cmod (f y - f x) < e/2" unfolding continuous_at Lim_at dist_norm using e by (drule_tac x="e/2" in spec) force obtain d2 where d2: "d2>0" "ball x d2 \ S" using \open S\ x by (auto simp: open_contains_ball) have dpos: "min d1 d2 > 0" using d1 d2 by simp { fix y assume yx: "y \ x" and close: "cmod (y - x) < min d1 d2" have y: "y \ S" using d2 close by (force simp: dist_norm norm_minus_commute) have "closed_segment x y \ S" using close d2 by (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1)) then have fxy: "f contour_integrable_on linepath x y" by (metis contour_integrable_continuous_linepath continuous_on_subset [OF contf]) then obtain i where i: "(f has_contour_integral i) (linepath x y)" by (auto simp: contour_integrable_on_def) then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" proof (rule has_contour_integral_bound_linepath) show "\u. u \ closed_segment x y \ cmod (f u - f x) \ e / 2" by (meson close d1_less le_less_trans less_imp_le min.strict_boundedE segment_bound1) qed (use e in simp) also have "\ < e * cmod (y - x)" by (simp add: e yx) finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using i yx by (simp add: contour_integral_unique divide_less_eq) } then have "\d>0. \y. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using dpos by blast } then have *: "(\y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \x\ 0" by (simp add: Lim_at dist_norm inverse_eq_divide) show ?thesis apply (simp add: has_field_derivative_def has_derivative_at2 bounded_linear_mult_right) apply (rule Lim_transform [OF * tendsto_eventually]) using \open S\ x apply (force simp: dist_norm open_contains_ball inverse_eq_divide [symmetric] eventually_at) done qed (** Existence of a primitive.*) lemma holomorphic_starlike_primitive: fixes f :: "complex \ complex" assumes contf: "continuous_on S f" and S: "starlike S" and os: "open S" and k: "finite k" and fcd: "\x. x \ S - k \ f field_differentiable at x" shows "\g. \x \ S. (g has_field_derivative f x) (at x)" proof - obtain a where a: "a\S" and a_cs: "\x. x\S \ closed_segment a x \ S" using S by (auto simp: starlike_def) { fix x b c assume "x \ S" "closed_segment b c \ S" then have abcs: "convex hull {a, b, c} \ S" by (simp add: a a_cs starlike_convex_subset) then have "continuous_on (convex hull {a, b, c}) f" by (simp add: continuous_on_subset [OF contf]) then have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" using abcs interior_subset by (force intro: fcd Cauchy_theorem_triangle_cofinite [OF _ k]) } note 0 = this show ?thesis apply (intro exI ballI) apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption) apply (metis a_cs) apply (metis has_chain_integral_chain_integral3 0) done qed lemma Cauchy_theorem_starlike: "\open S; starlike S; finite k; continuous_on S f; \x. x \ S - k \ f field_differentiable at x; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open) lemma Cauchy_theorem_starlike_simple: "\open S; starlike S; f holomorphic_on S; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI]) apply (simp_all add: holomorphic_on_imp_continuous_on) apply (metis at_within_open holomorphic_on_def) done subsection\Cauchy's theorem for a convex set\ text\For a convex set we can avoid assuming openness and boundary analyticity\ lemma triangle_contour_integrals_convex_primitive: assumes contf: "continuous_on S f" and S: "a \ S" "convex S" and x: "x \ S" and zer: "\b c. \b \ S; c \ S\ \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within S)" proof - let ?pathint = "\x y. contour_integral(linepath x y) f" { fix y assume y: "y \ S" have cont_ayf: "continuous_on (closed_segment a y) f" using S y by (meson contf continuous_on_subset convex_contains_segment) have xys: "closed_segment x y \ S" (*?*) using convex_contains_segment S x y by auto have "?pathint a y - ?pathint a x = ?pathint x y" using zer [OF x y] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force } note [simp] = this { fix e::real assume e: "0 < e" have cont_atx: "continuous (at x within S) f" using x S contf by (simp add: continuous_on_eq_continuous_within) then obtain d1 where d1: "d1>0" and d1_less: "\y. \y \ S; cmod (y - x) < d1\ \ cmod (f y - f x) < e/2" unfolding continuous_within Lim_within dist_norm using e by (drule_tac x="e/2" in spec) force { fix y assume yx: "y \ x" and close: "cmod (y - x) < d1" and y: "y \ S" have fxy: "f contour_integrable_on linepath x y" using convex_contains_segment S x y by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) then obtain i where i: "(f has_contour_integral i) (linepath x y)" by (auto simp: contour_integrable_on_def) then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" proof (rule has_contour_integral_bound_linepath) show "\u. u \ closed_segment x y \ cmod (f u - f x) \ e / 2" by (meson assms(3) close convex_contains_segment d1_less le_less_trans less_imp_le segment_bound1 subset_iff x y) qed (use e in simp) also have "\ < e * cmod (y - x)" by (simp add: e yx) finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using i yx by (simp add: contour_integral_unique divide_less_eq) } then have "\d>0. \y\S. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using d1 by blast } then have *: "((\y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \ 0) (at x within S)" by (simp add: Lim_within dist_norm inverse_eq_divide) show ?thesis apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) apply (rule Lim_transform [OF * tendsto_eventually]) using linordered_field_no_ub apply (force simp: inverse_eq_divide [symmetric] eventually_at) done qed lemma contour_integral_convex_primitive: assumes "convex S" "continuous_on S f" "\a b c. \a \ S; b \ S; c \ S\ \ (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (cases "S={}") case False with assms that show ?thesis by (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3) qed auto lemma holomorphic_convex_primitive: fixes f :: "complex \ complex" assumes "convex S" "finite K" and contf: "continuous_on S f" and fd: "\x. x \ interior S - K \ f field_differentiable at x" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (rule contour_integral_convex_primitive [OF \convex S\ contf Cauchy_theorem_triangle_cofinite]) have *: "convex hull {a, b, c} \ S" if "a \ S" "b \ S" "c \ S" for a b c by (simp add: \convex S\ hull_minimal that) show "continuous_on (convex hull {a, b, c}) f" if "a \ S" "b \ S" "c \ S" for a b c by (meson "*" contf continuous_on_subset that) show "f field_differentiable at x" if "a \ S" "b \ S" "c \ S" "x \ interior (convex hull {a, b, c}) - K" for a b c x by (metis "*" DiffD1 DiffD2 DiffI fd interior_mono subsetCE that) qed (use assms in \force+\) lemma holomorphic_convex_primitive': fixes f :: "complex \ complex" assumes "convex S" and "open S" and "f holomorphic_on S" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (rule holomorphic_convex_primitive) fix x assume "x \ interior S - {}" with assms show "f field_differentiable at x" by (auto intro!: holomorphic_on_imp_differentiable_at simp: interior_open) qed (use assms in \auto intro: holomorphic_on_imp_continuous_on\) corollary\<^marker>\tag unimportant\ Cauchy_theorem_convex: "\continuous_on S f; convex S; finite K; \x. x \ interior S - K \ f field_differentiable at x; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (metis holomorphic_convex_primitive Cauchy_theorem_primitive) corollary Cauchy_theorem_convex_simple: "\f holomorphic_on S; convex S; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" apply (rule Cauchy_theorem_convex [where K = "{}"]) apply (simp_all add: holomorphic_on_imp_continuous_on) using at_within_interior holomorphic_on_def interior_subset by fastforce text\In particular for a disc\ corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc: "\finite K; continuous_on (cball a e) f; \x. x \ ball a e - K \ f field_differentiable at x; valid_path g; path_image g \ cball a e; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (auto intro: Cauchy_theorem_convex) corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc_simple: "\f holomorphic_on (ball a e); valid_path g; path_image g \ ball a e; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (simp add: Cauchy_theorem_convex_simple) subsection\<^marker>\tag unimportant\ \Generalize integrability to local primitives\ lemma contour_integral_local_primitive_lemma: fixes f :: "complex\complex" shows "\g piecewise_differentiable_on {a..b}; \x. x \ s \ (f has_field_derivative f' x) (at x within s); \x. x \ {a..b} \ g x \ s\ \ (\x. f' (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b}" apply (cases "cbox a b = {}", force) apply (simp add: integrable_on_def) apply (rule exI) apply (rule contour_integral_primitive_lemma, assumption+) using atLeastAtMost_iff by blast lemma contour_integral_local_primitive_any: fixes f :: "complex \ complex" assumes gpd: "g piecewise_differentiable_on {a..b}" and dh: "\x. x \ s \ \d h. 0 < d \ (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" and gs: "\x. x \ {a..b} \ g x \ s" shows "(\x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}" proof - { fix x assume x: "a \ x" "x \ b" obtain d h where d: "0 < d" and h: "(\y. norm(y - g x) < d \ (h has_field_derivative f y) (at y within s))" using x gs dh by (metis atLeastAtMost_iff) have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast then obtain e where e: "e>0" and lessd: "\x'. x' \ {a..b} \ \x' - x\ < e \ cmod (g x' - g x) < d" using x d apply (auto simp: dist_norm continuous_on_iff) apply (drule_tac x=x in bspec) using x apply simp apply (drule_tac x=d in spec, auto) done have "\d>0. \u v. u \ x \ x \ v \ {u..v} \ ball x d \ (u \ v \ a \ u \ v \ b) \ (\x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}" apply (rule_tac x=e in exI) using e apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify) apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma) apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset) apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force) done } then show ?thesis by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified]) qed lemma contour_integral_local_primitive: fixes f :: "complex \ complex" assumes g: "valid_path g" "path_image g \ s" and dh: "\x. x \ s \ \d h. 0 < d \ (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" shows "f contour_integrable_on g" using g apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def has_integral_localized_vector_derivative integrable_on_def [symmetric]) using contour_integral_local_primitive_any [OF _ dh] by (meson image_subset_iff piecewise_C1_imp_differentiable) text\In particular if a function is holomorphic\ lemma contour_integrable_holomorphic: assumes contf: "continuous_on s f" and os: "open s" and k: "finite k" and g: "valid_path g" "path_image g \ s" and fcd: "\x. x \ s - k \ f field_differentiable at x" shows "f contour_integrable_on g" proof - { fix z assume z: "z \ s" obtain d where "d>0" and d: "ball z d \ s" using \open s\ z by (auto simp: open_contains_ball) then have contfb: "continuous_on (ball z d) f" using contf continuous_on_subset by blast obtain h where "\y\ball z d. (h has_field_derivative f y) (at y within ball z d)" by (metis holomorphic_convex_primitive [OF convex_ball k contfb fcd] d interior_subset Diff_iff subsetD) then have "\y\ball z d. (h has_field_derivative f y) (at y within s)" by (metis open_ball at_within_open d os subsetCE) then have "\h. (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" by (force simp: dist_norm norm_minus_commute) then have "\d h. 0 < d \ (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" using \0 < d\ by blast } then show ?thesis by (rule contour_integral_local_primitive [OF g]) qed lemma contour_integrable_holomorphic_simple: assumes fh: "f holomorphic_on S" and os: "open S" and g: "valid_path g" "path_image g \ S" shows "f contour_integrable_on g" apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g]) apply (simp add: fh holomorphic_on_imp_continuous_on) using fh by (simp add: field_differentiable_def holomorphic_on_open os) lemma continuous_on_inversediff: fixes z:: "'a::real_normed_field" shows "z \ S \ continuous_on S (\w. 1 / (w - z))" by (rule continuous_intros | force)+ lemma contour_integrable_inversediff: "\valid_path g; z \ path_image g\ \ (\w. 1 / (w-z)) contour_integrable_on g" apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"]) apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros) done text\Key fact that path integral is the same for a "nearby" path. This is the main lemma for the homotopy form of Cauchy's theorem and is also useful if we want "without loss of generality" to assume some nice properties of a path (e.g. smoothness). It can also be used to define the integrals of analytic functions over arbitrary continuous paths. This is just done for winding numbers now. \ text\A technical definition to avoid duplication of similar proofs, for paths joined at the ends versus looping paths\ definition linked_paths :: "bool \ (real \ 'a) \ (real \ 'a::topological_space) \ bool" where "linked_paths atends g h == (if atends then pathstart h = pathstart g \ pathfinish h = pathfinish g else pathfinish g = pathstart g \ pathfinish h = pathstart h)" text\This formulation covers two cases: \<^term>\g\ and \<^term>\h\ share their start and end points; \<^term>\g\ and \<^term>\h\ both loop upon themselves.\ lemma contour_integral_nearby: assumes os: "open S" and p: "path p" "path_image p \ S" shows "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ linked_paths atends g h \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" proof - have "\z. \e. z \ path_image p \ 0 < e \ ball z e \ S" using open_contains_ball os p(2) by blast then obtain ee where ee: "\z. z \ path_image p \ 0 < ee z \ ball z (ee z) \ S" by metis define cover where "cover = (\z. ball z (ee z/3)) ` (path_image p)" have "compact (path_image p)" by (metis p(1) compact_path_image) moreover have "path_image p \ (\c\path_image p. ball c (ee c / 3))" using ee by auto ultimately have "\D \ cover. finite D \ path_image p \ \D" by (simp add: compact_eq_Heine_Borel cover_def) then obtain D where D: "D \ cover" "finite D" "path_image p \ \D" by blast then obtain k where k: "k \ {0..1}" "finite k" and D_eq: "D = ((\z. ball z (ee z / 3)) \ p) ` k" apply (simp add: cover_def path_image_def image_comp) apply (blast dest!: finite_subset_image [OF \finite D\]) done then have kne: "k \ {}" using D by auto have pi: "\i. i \ k \ p i \ path_image p" using k by (auto simp: path_image_def) then have eepi: "\i. i \ k \ 0 < ee((p i))" by (metis ee) define e where "e = Min((ee \ p) ` k)" have fin_eep: "finite ((ee \ p) ` k)" using k by blast have "0 < e" using ee k by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi) have "uniformly_continuous_on {0..1} p" using p by (simp add: path_def compact_uniformly_continuous) then obtain d::real where d: "d>0" and de: "\x x'. \x' - x\ < d \ x\{0..1} \ x'\{0..1} \ cmod (p x' - p x) < e/3" unfolding uniformly_continuous_on_def dist_norm real_norm_def by (metis divide_pos_pos \0 < e\ zero_less_numeral) then obtain N::nat where N: "N>0" "inverse N < d" using real_arch_inverse [of d] by auto show ?thesis proof (intro exI conjI allI; clarify?) show "e/3 > 0" using \0 < e\ by simp fix g h assume g: "valid_path g" and ghp: "\t\{0..1}. cmod (g t - p t) < e / 3 \ cmod (h t - p t) < e / 3" and h: "valid_path h" and joins: "linked_paths atends g h" { fix t::real assume t: "0 \ t" "t \ 1" then obtain u where u: "u \ k" and ptu: "p t \ ball(p u) (ee(p u) / 3)" using \path_image p \ \D\ D_eq by (force simp: path_image_def) then have ele: "e \ ee (p u)" using fin_eep by (simp add: e_def) have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3" using ghp t by auto with ele have "cmod (g t - p t) < ee (p u) / 3" "cmod (h t - p t) < ee (p u) / 3" by linarith+ then have "g t \ ball(p u) (ee(p u))" "h t \ ball(p u) (ee(p u))" using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"] norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u by (force simp: dist_norm ball_def norm_minus_commute)+ then have "g t \ S" "h t \ S" using ee u k by (auto simp: path_image_def ball_def) } then have ghs: "path_image g \ S" "path_image h \ S" by (auto simp: path_image_def) moreover { fix f assume fhols: "f holomorphic_on S" then have fpa: "f contour_integrable_on g" "f contour_integrable_on h" using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple by blast+ have contf: "continuous_on S f" by (simp add: fhols holomorphic_on_imp_continuous_on) { fix z assume z: "z \ path_image p" have "f holomorphic_on ball z (ee z)" using fhols ee z holomorphic_on_subset by blast then have "\ff. (\w \ ball z (ee z). (ff has_field_derivative f w) (at w))" using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified] by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball) } then obtain ff where ff: "\z w. \z \ path_image p; w \ ball z (ee z)\ \ (ff z has_field_derivative f w) (at w)" by metis { fix n assume n: "n \ N" then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f = contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f" proof (induct n) case 0 show ?case by simp next case (Suc n) obtain t where t: "t \ k" and "p (n/N) \ ball(p t) (ee(p t) / 3)" using \path_image p \ \D\ [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems by (force simp: path_image_def) then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3" by (simp add: dist_norm) have e3le: "e/3 \ ee (p t) / 3" using fin_eep t by (simp add: e_def) { fix x assume x: "n/N \ x" "x \ (1 + n)/N" then have nN01: "0 \ n/N" "(1 + n)/N \ 1" using Suc.prems by auto then have x01: "0 \ x" "x \ 1" using x by linarith+ have "cmod (p t - p x) < ee (p t) / 3 + e/3" proof (rule norm_diff_triangle_less [OF ptu de]) show "\real n / real N - x\ < d" using x N by (auto simp: field_simps) qed (use x01 Suc.prems in auto) then have ptx: "cmod (p t - p x) < 2*ee (p t)/3" using e3le eepi [OF t] by simp have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 " apply (rule norm_diff_triangle_less [OF ptx]) using ghp x01 by (simp add: norm_minus_commute) also have "\ \ ee (p t)" using e3le eepi [OF t] by simp finally have gg: "cmod (p t - g x) < ee (p t)" . have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 " apply (rule norm_diff_triangle_less [OF ptx]) using ghp x01 by (simp add: norm_minus_commute) also have "\ \ ee (p t)" using e3le eepi [OF t] by simp finally have "cmod (p t - g x) < ee (p t)" "cmod (p t - h x) < ee (p t)" using gg by auto } note ptgh_ee = this have "closed_segment (g (real n / real N)) (h (real n / real N)) = path_image (linepath (h (n/N)) (g (n/N)))" by (simp add: closed_segment_commute) also have pi_hgn: "\ \ ball (p t) (ee (p t))" using ptgh_ee [of "n/N"] Suc.prems by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) finally have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \ S" using ee pi t by blast have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N))) \ ball (p t) (ee (p t))" using ptgh_ee [of "(1+n)/N"] Suc.prems by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \ S" using \N>0\ Suc.prems ee pi t by (auto simp: Path_Connected.path_image_join field_simps) have pi_subset_ball: "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++ subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N))) \ ball (p t) (ee (p t))" apply (intro subset_path_image_join pi_hgn pi_ghn') using \N>0\ Suc.prems apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) done have pi0: "(f has_contour_integral 0) (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++ subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))" apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"]) apply (metis ff open_ball at_within_open pi t) using Suc.prems pi_subset_ball apply (simp_all add: valid_path_join valid_path_subpath g h) done have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g" using Suc.prems by (simp add: contour_integrable_subpath g fpa) have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))" using gh_n's by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))" using gh_ns by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f + contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f + contour_integral (subpath ((Suc n) / N) (n/N) h) f + contour_integral (linepath (h (n/N)) (g (n/N))) f = 0" using contour_integral_unique [OF pi0] Suc.prems by (simp add: g h fpa valid_path_subpath contour_integrable_subpath fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc) have *: "\hn he hn' gn gd gn' hgn ghn gh0 ghn'. \hn - gn = ghn - gh0; gd + ghn' + he + hgn = (0::complex); hn - he = hn'; gn + gd = gn'; hgn = -ghn\ \ hn' - gn' = ghn' - gh0" by (auto simp: algebra_simps) have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f" unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"] using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath) also have "\ = contour_integral (subpath 0 ((Suc n) / N) h) f" using Suc.prems by (simp add: contour_integral_subpath_combine h fpa) finally have pi0_eq: "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = contour_integral (subpath 0 ((Suc n) / N) h) f" . show ?case apply (rule * [OF Suc.hyps eq0 pi0_eq]) using Suc.prems apply (simp_all add: g h fpa contour_integral_subpath_combine contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath continuous_on_subset [OF contf gh_ns]) done qed } note ind = this have "contour_integral h f = contour_integral g f" using ind [OF order_refl] N joins by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm) } ultimately show "path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f)" by metis qed qed - lemma assumes "open S" "path p" "path_image p \ S" shows contour_integral_nearby_ends: "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ pathstart h = pathstart g \ pathfinish h = pathfinish g \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" and contour_integral_nearby_loops: "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ pathfinish g = pathstart g \ pathfinish h = pathstart h \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" using contour_integral_nearby [OF assms, where atends=True] using contour_integral_nearby [OF assms, where atends=False] unfolding linked_paths_def by simp_all lemma C1_differentiable_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ p C1_differentiable_on S" by (metis continuous_on_polymonial_function C1_differentiable_on_def has_vector_derivative_polynomial_function) lemma valid_path_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ valid_path p" by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) lemma valid_path_subpath_trivial [simp]: fixes g :: "real \ 'a::euclidean_space" shows "z \ g x \ valid_path (subpath x x g)" by (simp add: subpath_def valid_path_polynomial_function) lemma contour_integral_bound_exists: assumes S: "open S" and g: "valid_path g" and pag: "path_image g \ S" shows "\L. 0 < L \ (\f B. f holomorphic_on S \ (\z \ S. norm(f z) \ B) \ norm(contour_integral g f) \ L*B)" proof - have "path g" using g by (simp add: valid_path_imp_path) then obtain d::real and p where d: "0 < d" and p: "polynomial_function p" "path_image p \ S" and pi: "\f. f holomorphic_on S \ contour_integral g f = contour_integral p f" using contour_integral_nearby_ends [OF S \path g\ pag] apply clarify apply (drule_tac x=g in spec) apply (simp only: assms) apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function) done then obtain p' where p': "polynomial_function p'" "\x. (p has_vector_derivative (p' x)) (at x)" by (blast intro: has_vector_derivative_polynomial_function that) then have "bounded(p' ` {0..1})" using continuous_on_polymonial_function by (force simp: intro!: compact_imp_bounded compact_continuous_image) then obtain L where L: "L>0" and nop': "\x. \0 \ x; x \ 1\ \ norm (p' x) \ L" by (force simp: bounded_pos) { fix f B assume f: "f holomorphic_on S" and B: "\z. z\S \ cmod (f z) \ B" then have "f contour_integrable_on p \ valid_path p" using p S by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on) moreover have "cmod (vector_derivative p (at x)) * cmod (f (p x)) \ L * B" if "0 \ x" "x \ 1" for x proof (rule mult_mono) show "cmod (vector_derivative p (at x)) \ L" by (metis nop' p'(2) that vector_derivative_at) show "cmod (f (p x)) \ B" by (metis B atLeastAtMost_iff imageI p(2) path_defs(4) subset_eq that) qed (use \L>0\ in auto) ultimately have "cmod (contour_integral g f) \ L * B" apply (simp only: pi [OF f]) apply (simp only: contour_integral_integral) apply (rule order_trans [OF integral_norm_bound_integral]) apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult) done } then show ?thesis by (force simp: L contour_integral_integral) qed -text\We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\ - -subsection \Winding Numbers\ - -definition\<^marker>\tag important\ winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where - "winding_number_prop \ z e p n \ - valid_path p \ z \ path_image p \ - pathstart p = pathstart \ \ - pathfinish p = pathfinish \ \ - (\t \ {0..1}. norm(\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - -definition\<^marker>\tag important\ winding_number:: "[real \ complex, complex] \ complex" where - "winding_number \ z \ SOME n. \e > 0. \p. winding_number_prop \ z e p n" - - -lemma winding_number: - assumes "path \" "z \ path_image \" "0 < e" - shows "\p. winding_number_prop \ z e p (winding_number \ z)" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain d - where d: "d>0" - and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ - (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ - pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ - path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ - (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" - using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) - then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ - (\t \ {0..1}. norm(h t - \ t) < d/2)" - using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto - define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" - have "\n. \e > 0. \p. winding_number_prop \ z e p n" - proof (rule_tac x=nn in exI, clarify) - fix e::real - assume e: "e>0" - obtain p where p: "polynomial_function p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d/2))" - using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by auto - have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" - by (auto simp: intro!: holomorphic_intros) - then show "\p. winding_number_prop \ z e p nn" - apply (rule_tac x=p in exI) - using pi_eq [of h p] h p d - apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def winding_number_prop_def) - done - qed - then show ?thesis - unfolding winding_number_def by (rule someI2_ex) (blast intro: \0) -qed - -lemma winding_number_unique: - assumes \: "path \" "z \ path_image \" - and pi: "\e. e>0 \ \p. winding_number_prop \ z e p n" - shows "winding_number \ z = n" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain e - where e: "e>0" - and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); - pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ - contour_integral h2 f = contour_integral h1 f" - using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) - obtain p where p: "winding_number_prop \ z e p n" - using pi [OF e] by blast - obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" - using winding_number [OF \ e] by blast - have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" - using p by (auto simp: winding_number_prop_def) - also have "\ = contour_integral q (\w. 1 / (w - z))" - proof (rule pi_eq) - show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" - by (auto intro!: holomorphic_intros) - qed (use p q in \auto simp: winding_number_prop_def norm_minus_commute\) - also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" - using q by (auto simp: winding_number_prop_def) - finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . - then show ?thesis - by simp -qed - -(*NB not winding_number_prop here due to the loop in p*) -lemma winding_number_unique_loop: - assumes \: "path \" "z \ path_image \" - and loop: "pathfinish \ = pathstart \" - and pi: - "\e. e>0 \ \p. valid_path p \ z \ path_image p \ - pathfinish p = pathstart p \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - shows "winding_number \ z = n" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain e - where e: "e>0" - and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); - pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ - contour_integral h2 f = contour_integral h1 f" - using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) - obtain p where p: - "valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - using pi [OF e] by blast - obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" - using winding_number [OF \ e] by blast - have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" - using p by auto - also have "\ = contour_integral q (\w. 1 / (w - z))" - proof (rule pi_eq) - show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" - by (auto intro!: holomorphic_intros) - qed (use p q loop in \auto simp: winding_number_prop_def norm_minus_commute\) - also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" - using q by (auto simp: winding_number_prop_def) - finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . - then show ?thesis - by simp -qed - -proposition winding_number_valid_path: - assumes "valid_path \" "z \ path_image \" - shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" - by (rule winding_number_unique) - (use assms in \auto simp: valid_path_imp_path winding_number_prop_def\) - -proposition has_contour_integral_winding_number: - assumes \: "valid_path \" "z \ path_image \" - shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" -by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) - -lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" - by (simp add: winding_number_valid_path) - -lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" - by (simp add: path_image_subpath winding_number_valid_path) - -lemma winding_number_join: - assumes \1: "path \1" "z \ path_image \1" - and \2: "path \2" "z \ path_image \2" - and "pathfinish \1 = pathstart \2" - shows "winding_number(\1 +++ \2) z = winding_number \1 z + winding_number \2 z" -proof (rule winding_number_unique) - show "\p. winding_number_prop (\1 +++ \2) z e p - (winding_number \1 z + winding_number \2 z)" if "e > 0" for e - proof - - obtain p1 where "winding_number_prop \1 z e p1 (winding_number \1 z)" - using \0 < e\ \1 winding_number by blast - moreover - obtain p2 where "winding_number_prop \2 z e p2 (winding_number \2 z)" - using \0 < e\ \2 winding_number by blast - ultimately - have "winding_number_prop (\1+++\2) z e (p1+++p2) (winding_number \1 z + winding_number \2 z)" - using assms - apply (simp add: winding_number_prop_def not_in_path_image_join contour_integrable_inversediff algebra_simps) - apply (auto simp: joinpaths_def) - done - then show ?thesis - by blast - qed -qed (use assms in \auto simp: not_in_path_image_join\) - -lemma winding_number_reversepath: - assumes "path \" "z \ path_image \" - shows "winding_number(reversepath \) z = - (winding_number \ z)" -proof (rule winding_number_unique) - show "\p. winding_number_prop (reversepath \) z e p (- winding_number \ z)" if "e > 0" for e - proof - - obtain p where "winding_number_prop \ z e p (winding_number \ z)" - using \0 < e\ assms winding_number by blast - then have "winding_number_prop (reversepath \) z e (reversepath p) (- winding_number \ z)" - using assms - apply (simp add: winding_number_prop_def contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) - apply (auto simp: reversepath_def) - done - then show ?thesis - by blast - qed -qed (use assms in auto) - -lemma winding_number_shiftpath: - assumes \: "path \" "z \ path_image \" - and "pathfinish \ = pathstart \" "a \ {0..1}" - shows "winding_number(shiftpath a \) z = winding_number \ z" -proof (rule winding_number_unique_loop) - show "\p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ - (\t\{0..1}. cmod (shiftpath a \ t - p t) < e) \ - contour_integral p (\w. 1 / (w - z)) = - complex_of_real (2 * pi) * \ * winding_number \ z" - if "e > 0" for e - proof - - obtain p where "winding_number_prop \ z e p (winding_number \ z)" - using \0 < e\ assms winding_number by blast - then show ?thesis - apply (rule_tac x="shiftpath a p" in exI) - using assms that - apply (auto simp: winding_number_prop_def path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath contour_integral_shiftpath) - apply (simp add: shiftpath_def) - done - qed -qed (use assms in \auto simp: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath\) - -lemma winding_number_split_linepath: - assumes "c \ closed_segment a b" "z \ closed_segment a b" - shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" -proof - - have "z \ closed_segment a c" "z \ closed_segment c b" - using assms by (meson convex_contains_segment convex_segment ends_in_segment subsetCE)+ - then show ?thesis - using assms - by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) -qed - -lemma winding_number_cong: - "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" - by (simp add: winding_number_def winding_number_prop_def pathstart_def pathfinish_def) - -lemma winding_number_constI: - assumes "c\z" "\t. \0\t; t\1\ \ g t = c" - shows "winding_number g z = 0" -proof - - have "winding_number g z = winding_number (linepath c c) z" - apply (rule winding_number_cong) - using assms unfolding linepath_def by auto - moreover have "winding_number (linepath c c) z =0" - apply (rule winding_number_trivial) - using assms by auto - ultimately show ?thesis by auto -qed - -lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" - unfolding winding_number_def -proof (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) - fix n e g - assume "0 < e" and g: "winding_number_prop p z e g n" - then show "\r. winding_number_prop (\w. p w - z) 0 e r n" - by (rule_tac x="\t. g t - z" in exI) - (force simp: winding_number_prop_def contour_integral_integral valid_path_def path_defs - vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) -next - fix n e g - assume "0 < e" and g: "winding_number_prop (\w. p w - z) 0 e g n" - then show "\r. winding_number_prop p z e r n" - apply (rule_tac x="\t. g t + z" in exI) - apply (simp add: winding_number_prop_def contour_integral_integral valid_path_def path_defs - piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) - apply (force simp: algebra_simps) - done -qed - -subsubsection\<^marker>\tag unimportant\ \Some lemmas about negating a path\ - -lemma valid_path_negatepath: "valid_path \ \ valid_path (uminus \ \)" - unfolding o_def using piecewise_C1_differentiable_neg valid_path_def by blast - -lemma has_contour_integral_negatepath: - assumes \: "valid_path \" and cint: "((\z. f (- z)) has_contour_integral - i) \" - shows "(f has_contour_integral i) (uminus \ \)" -proof - - obtain S where cont: "continuous_on {0..1} \" and "finite S" and diff: "\ C1_differentiable_on {0..1} - S" - using \ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - have "((\x. - (f (- \ x) * vector_derivative \ (at x within {0..1}))) has_integral i) {0..1}" - using cint by (auto simp: has_contour_integral_def dest: has_integral_neg) - then - have "((\x. f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1})) has_integral i) {0..1}" - proof (rule rev_iffD1 [OF _ has_integral_spike_eq]) - show "negligible S" - by (simp add: \finite S\ negligible_finite) - show "f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1}) = - - (f (- \ x) * vector_derivative \ (at x within {0..1}))" - if "x \ {0..1} - S" for x - proof - - have "vector_derivative (uminus \ \) (at x within cbox 0 1) = - vector_derivative \ (at x within cbox 0 1)" - proof (rule vector_derivative_within_cbox) - show "(uminus \ \ has_vector_derivative - vector_derivative \ (at x within cbox 0 1)) (at x within cbox 0 1)" - using that unfolding o_def - by (metis C1_differentiable_on_eq UNIV_I diff differentiable_subset has_vector_derivative_minus subsetI that vector_derivative_works) - qed (use that in auto) - then show ?thesis - by simp - qed - qed - then show ?thesis by (simp add: has_contour_integral_def) -qed - -lemma winding_number_negatepath: - assumes \: "valid_path \" and 0: "0 \ path_image \" - shows "winding_number(uminus \ \) 0 = winding_number \ 0" -proof - - have "(/) 1 contour_integrable_on \" - using "0" \ contour_integrable_inversediff by fastforce - then have "((\z. 1/z) has_contour_integral contour_integral \ ((/) 1)) \" - by (rule has_contour_integral_integral) - then have "((\z. 1 / - z) has_contour_integral - contour_integral \ ((/) 1)) \" - using has_contour_integral_neg by auto - then show ?thesis - using assms - apply (simp add: winding_number_valid_path valid_path_negatepath image_def path_defs) - apply (simp add: contour_integral_unique has_contour_integral_negatepath) - done -qed - -lemma contour_integrable_negatepath: - assumes \: "valid_path \" and pi: "(\z. f (- z)) contour_integrable_on \" - shows "f contour_integrable_on (uminus \ \)" - by (metis \ add.inverse_inverse contour_integrable_on_def has_contour_integral_negatepath pi) - -(* A combined theorem deducing several things piecewise.*) -lemma winding_number_join_pos_combined: - "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); - valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ - \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" - by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) - - -subsubsection\<^marker>\tag unimportant\ \Useful sufficient conditions for the winding number to be positive\ - -lemma Re_winding_number: - "\valid_path \; z \ path_image \\ - \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" -by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) - -lemma winding_number_pos_le: - assumes \: "valid_path \" "z \ path_image \" - and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" - shows "0 \ Re(winding_number \ z)" -proof - - have ge0: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x - using ge by (simp add: Complex.Im_divide algebra_simps x) - let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" - let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" - have hi: "(?vd has_integral ?int z) (cbox 0 1)" - unfolding box_real - apply (subst has_contour_integral [symmetric]) - using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) - have "0 \ Im (?int z)" - proof (rule has_integral_component_nonneg [of \, simplified]) - show "\x. x \ cbox 0 1 \ 0 \ Im (if 0 < x \ x < 1 then ?vd x else 0)" - by (force simp: ge0) - show "((\x. if 0 < x \ x < 1 then ?vd x else 0) has_integral ?int z) (cbox 0 1)" - by (rule has_integral_spike_interior [OF hi]) simp - qed - then show ?thesis - by (simp add: Re_winding_number [OF \] field_simps) -qed - -lemma winding_number_pos_lt_lemma: - assumes \: "valid_path \" "z \ path_image \" - and e: "0 < e" - and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" - shows "0 < Re(winding_number \ z)" -proof - - let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" - let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" - have hi: "(?vd has_integral ?int z) (cbox 0 1)" - unfolding box_real - apply (subst has_contour_integral [symmetric]) - using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) - have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" - proof (rule has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified]) - show "((\x. if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e) has_integral ?int z) {0..1}" - by (rule has_integral_spike_interior [OF hi, simplified box_real]) (use e in simp) - show "\x. 0 \ x \ x \ 1 \ - e \ Im (if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e)" - by (simp add: ge) - qed (use has_integral_const_real [of _ 0 1] in auto) - with e show ?thesis - by (simp add: Re_winding_number [OF \] field_simps) -qed - -lemma winding_number_pos_lt: - assumes \: "valid_path \" "z \ path_image \" - and e: "0 < e" - and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" - shows "0 < Re (winding_number \ z)" -proof - - have bm: "bounded ((\w. w - z) ` (path_image \))" - using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) - then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" - using bounded_pos [THEN iffD1, OF bm] by blast - { fix x::real assume x: "0 < x" "x < 1" - then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] - by (simp add: path_image_def power2_eq_square mult_mono') - with x have "\ x \ z" using \ - using path_image_def by fastforce - then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" - using B ge [OF x] B2 e - apply (rule_tac y="e / (cmod (\ x - z))\<^sup>2" in order_trans) - apply (auto simp: divide_left_mono divide_right_mono) - done - then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" - by (simp add: complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) - } note * = this - show ?thesis - using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) -qed - -subsection\The winding number is an integer\ - -text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, - Also on page 134 of Serge Lang's book with the name title, etc.\ - -lemma exp_fg: - fixes z::complex - assumes g: "(g has_vector_derivative g') (at x within s)" - and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" - and z: "g x \ z" - shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" -proof - - have *: "(exp \ (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" - using assms unfolding has_vector_derivative_def scaleR_conv_of_real - by (auto intro!: derivative_eq_intros) - show ?thesis - apply (rule has_vector_derivative_eq_rhs) - using z - apply (auto intro!: derivative_eq_intros * [unfolded o_def] g) - done -qed - -lemma winding_number_exp_integral: - fixes z::complex - assumes \: "\ piecewise_C1_differentiable_on {a..b}" - and ab: "a \ b" - and z: "z \ \ ` {a..b}" - shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" - (is "?thesis1") - "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" - (is "?thesis2") -proof - - let ?D\ = "\x. vector_derivative \ (at x)" - have [simp]: "\x. a \ x \ x \ b \ \ x \ z" - using z by force - have cong: "continuous_on {a..b} \" - using \ by (simp add: piecewise_C1_differentiable_on_def) - obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" - using \ by (force simp: piecewise_C1_differentiable_on_def) - have \: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) - moreover have "{a<.. {a..b} - k" - by force - ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" - by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def at_within_open) - { fix w - assume "w \ z" - have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" - by (auto simp: dist_norm intro!: continuous_intros) - moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" - by (auto simp: intro!: derivative_eq_intros) - ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" - using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] - by (force simp: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) - } - then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" - by meson - have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" - unfolding integrable_on_def [symmetric] - proof (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \]]) - show "\d h. 0 < d \ - (\y. cmod (y - w) < d \ (h has_field_derivative inverse (y - z))(at y within - {z}))" - if "w \ - {z}" for w - apply (rule_tac x="norm(w - z)" in exI) - using that inverse_eq_divide has_field_derivative_at_within h - by (metis Compl_insert DiffD2 insertCI right_minus_eq zero_less_norm_iff) - qed simp - have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" - unfolding box_real [symmetric] divide_inverse_commute - by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) - with ab show ?thesis1 - by (simp add: divide_inverse_commute integral_def integrable_on_def) - { fix t - assume t: "t \ {a..b}" - have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" - using z by (auto intro!: continuous_intros simp: dist_norm) - have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" - unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) - obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ - (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" - using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] - by simp (auto simp: ball_def dist_norm that) - { fix x D - assume x: "x \ k" "a < x" "x < b" - then have "x \ interior ({a..b} - k)" - using open_subset_interior [OF \] by fastforce - then have con: "isCont ?D\ x" - using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) - then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" - by (rule continuous_at_imp_continuous_within) - have gdx: "\ differentiable at x" - using x by (simp add: g_diff_at) - have "\d. \x \ k; a < x; x < b; - (\ has_vector_derivative d) (at x); a \ t; t \ b\ - \ ((\x. integral {a..x} - (\x. ?D\ x / - (\ x - z))) has_vector_derivative - d / (\ x - z)) - (at x within {a..b})" - apply (rule has_vector_derivative_eq_rhs) - apply (rule integral_has_vector_derivative_continuous_at [where S = "{}", simplified]) - apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+ - done - then have "((\c. exp (- integral {a..c} (\x. ?D\ x / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) - (at x within {a..b})" - using x gdx t - apply (clarsimp simp add: differentiable_iff_scaleR) - apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_withinI) - apply (simp_all add: has_vector_derivative_def [symmetric]) - done - } note * = this - have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" - apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) - using t - apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous_1 [OF vg_int] simp add: ab)+ - done - } - with ab show ?thesis2 - by (simp add: divide_inverse_commute integral_def) -qed - -lemma winding_number_exp_2pi: - "\path p; z \ path_image p\ - \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" -using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def winding_number_prop_def - by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) - -lemma integer_winding_number_eq: - assumes \: "path \" and z: "z \ path_image \" - shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" -proof - - obtain p where p: "valid_path p" "z \ path_image p" - "pathstart p = pathstart \" "pathfinish p = pathfinish \" - and eq: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF assms, of 1] unfolding winding_number_prop_def by auto - then have wneq: "winding_number \ z = winding_number p z" - using eq winding_number_valid_path by force - have iff: "(winding_number \ z \ \) \ (exp (contour_integral p (\w. 1 / (w - z))) = 1)" - using eq by (simp add: exp_eq_1 complex_is_Int_iff) - have "exp (contour_integral p (\w. 1 / (w - z))) = (\ 1 - z) / (\ 0 - z)" - using p winding_number_exp_integral(2) [of p 0 1 z] - apply (simp add: valid_path_def path_defs contour_integral_integral exp_minus field_split_simps) - by (metis path_image_def pathstart_def pathstart_in_path_image) - then have "winding_number p z \ \ \ pathfinish p = pathstart p" - using p wneq iff by (auto simp: path_defs) - then show ?thesis using p eq - by (auto simp: winding_number_valid_path) -qed - -theorem integer_winding_number: - "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" -by (metis integer_winding_number_eq) - - -text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) - We can thus bound the winding number of a path that doesn't intersect a given ray. \ - -lemma winding_number_pos_meets: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" - and w: "w \ z" - shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" -proof - - have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" - using z by (auto simp: path_image_def) - have [simp]: "z \ \ ` {0..1}" - using path_image_def z by auto - have gpd: "\ piecewise_C1_differentiable_on {0..1}" - using \ valid_path_def by blast - define r where "r = (w - z) / (\ 0 - z)" - have [simp]: "r \ 0" - using w z by (auto simp: r_def) - have cont: "continuous_on {0..1} - (\x. Im (integral {0..x} (\x. vector_derivative \ (at x) / (\ x - z))))" - by (intro continuous_intros indefinite_integral_continuous_1 winding_number_exp_integral [OF gpd]; simp) - have "Arg2pi r \ 2*pi" - by (simp add: Arg2pi less_eq_real_def) - also have "\ \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" - using 1 - apply (simp add: winding_number_valid_path [OF \ z] contour_integral_integral) - apply (simp add: Complex.Re_divide field_simps power2_eq_square) - done - finally have "Arg2pi r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . - then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" - by (simp add: Arg2pi_ge_0 cont IVT') - then obtain t where t: "t \ {0..1}" - and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" - by blast - define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" - have iArg: "Arg2pi r = Im i" - using eqArg by (simp add: i_def) - have gpdt: "\ piecewise_C1_differentiable_on {0..t}" - by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) - have "exp (- i) * (\ t - z) = \ 0 - z" - unfolding i_def - apply (rule winding_number_exp_integral [OF gpdt]) - using t z unfolding path_image_def by force+ - then have *: "\ t - z = exp i * (\ 0 - z)" - by (simp add: exp_minus field_simps) - then have "(w - z) = r * (\ 0 - z)" - by (simp add: r_def) - then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" - apply simp - apply (subst Complex_Transcendental.Arg2pi_eq [of r]) - apply (simp add: iArg) - using * apply (simp add: exp_eq_polar field_simps) - done - with t show ?thesis - by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) -qed - -lemma winding_number_big_meets: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" - and w: "w \ z" - shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" -proof - - { assume "Re (winding_number \ z) \ - 1" - then have "Re (winding_number (reversepath \) z) \ 1" - by (simp add: \ valid_path_imp_path winding_number_reversepath z) - moreover have "valid_path (reversepath \)" - using \ valid_path_imp_reverse by auto - moreover have "z \ path_image (reversepath \)" - by (simp add: z) - ultimately have "\a::real. 0 < a \ z + a*(w - z) \ path_image (reversepath \)" - using winding_number_pos_meets w by blast - then have ?thesis - by simp - } - then show ?thesis - using assms - by (simp add: abs_if winding_number_pos_meets split: if_split_asm) -qed - -lemma winding_number_less_1: - fixes z::complex - shows - "\valid_path \; z \ path_image \; w \ z; - \a::real. 0 < a \ z + a*(w - z) \ path_image \\ - \ Re(winding_number \ z) < 1" - by (auto simp: not_less dest: winding_number_big_meets) - -text\One way of proving that WN=1 for a loop.\ -lemma winding_number_eq_1: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" - and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" - shows "winding_number \ z = 1" -proof - - have "winding_number \ z \ Ints" - by (simp add: \ integer_winding_number loop valid_path_imp_path z) - then show ?thesis - using 0 2 by (auto simp: Ints_def) -qed - -subsection\Continuity of winding number and invariance on connected sets\ - -lemma continuous_at_winding_number: - fixes z::complex - assumes \: "path \" and z: "z \ path_image \" - shows "continuous (at z) (winding_number \)" -proof - - obtain e where "e>0" and cbg: "cball z e \ - path_image \" - using open_contains_cball [of "- path_image \"] z - by (force simp: closed_def [symmetric] closed_path_image [OF \]) - then have ppag: "path_image \ \ - cball z (e/2)" - by (force simp: cball_def dist_norm) - have oc: "open (- cball z (e / 2))" - by (simp add: closed_def [symmetric]) - obtain d where "d>0" and pi_eq: - "\h1 h2. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); - pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ - \ - path_image h1 \ - cball z (e / 2) \ - path_image h2 \ - cball z (e / 2) \ - (\f. f holomorphic_on - cball z (e / 2) \ contour_integral h2 f = contour_integral h1 f)" - using contour_integral_nearby_ends [OF oc \ ppag] by metis - obtain p where p: "valid_path p" "z \ path_image p" - "pathstart p = pathstart \ \ pathfinish p = pathfinish \" - and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e / 2" - and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF \ z, of "min d e / 2"] \d>0\ \e>0\ by (auto simp: winding_number_prop_def) - { fix w - assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" - then have wnotp: "w \ path_image p" - using cbg \d>0\ \e>0\ - apply (simp add: path_image_def cball_def dist_norm, clarify) - apply (frule pg) - apply (drule_tac c="\ x" in subsetD) - apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) - done - have wnotg: "w \ path_image \" - using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) - { fix k::real - assume k: "k>0" - then obtain q where q: "valid_path q" "w \ path_image q" - "pathstart q = pathstart \ \ pathfinish q = pathfinish \" - and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" - and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" - using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k - by (force simp: min_divide_distrib_right winding_number_prop_def) - have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" - apply (rule pi_eq [OF \valid_path q\ \valid_path p\, THEN conjunct2, THEN conjunct2, rule_format]) - apply (frule pg) - apply (frule qg) - using p q \d>0\ e2 - apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) - done - then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" - by (simp add: pi qi) - } note pip = this - have "path p" - using p by (simp add: valid_path_imp_path) - then have "winding_number p w = winding_number \ w" - apply (rule winding_number_unique [OF _ wnotp]) - apply (rule_tac x=p in exI) - apply (simp add: p wnotp min_divide_distrib_right pip winding_number_prop_def) - done - } note wnwn = this - obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" - using p open_contains_cball [of "- path_image p"] - by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) - obtain L - where "L>0" - and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); - \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ - cmod (contour_integral p f) \ L * B" - using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by force - { fix e::real and w::complex - assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" - then have [simp]: "w \ path_image p" - using cbp p(2) \0 < pe\ - by (force simp: dist_norm norm_minus_commute path_image_def cball_def) - have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = - contour_integral p (\x. 1/(x - w) - 1/(x - z))" - by (simp add: p contour_integrable_inversediff contour_integral_diff) - { fix x - assume pe: "3/4 * pe < cmod (z - x)" - have "cmod (w - x) < pe/4 + cmod (z - x)" - by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) - then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp - have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" - using norm_diff_triangle_le by blast - also have "\ < pe/4 + cmod (w - x)" - using w by (simp add: norm_minus_commute) - finally have "pe/2 < cmod (w - x)" - using pe by auto - then have "(pe/2)^2 < cmod (w - x) ^ 2" - apply (rule power_strict_mono) - using \pe>0\ by auto - then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" - by (simp add: power_divide) - have "8 * L * cmod (w - z) < e * pe\<^sup>2" - using w \L>0\ by (simp add: field_simps) - also have "\ < e * 4 * cmod (w - x) * cmod (w - x)" - using pe2 \e>0\ by (simp add: power2_eq_square) - also have "\ < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" - using wx - apply (rule mult_strict_left_mono) - using pe2 e not_less_iff_gr_or_eq by fastforce - finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" - by simp - also have "\ \ e * cmod (w - x) * cmod (z - x)" - using e by simp - finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . - have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" - apply (cases "x=z \ x=w") - using pe \pe>0\ w \L>0\ - apply (force simp: norm_minus_commute) - using wx w(2) \L>0\ pe pe2 Lwz - apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) - done - } note L_cmod_le = this - have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" - apply (rule L) - using \pe>0\ w - apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) - using \pe>0\ w \L>0\ - apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) - done - have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" - apply simp - apply (rule le_less_trans [OF *]) - using \L>0\ e - apply (force simp: field_simps) - done - then have "cmod (winding_number p w - winding_number p z) < e" - using pi_ge_two e - by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) - } note cmod_wn_diff = this - then have "isCont (winding_number p) z" - apply (simp add: continuous_at_eps_delta, clarify) - apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) - using \pe>0\ \L>0\ - apply (simp add: dist_norm cmod_wn_diff) - done - then show ?thesis - apply (rule continuous_transform_within [where d = "min d e / 2"]) - apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) - done -qed - -corollary continuous_on_winding_number: - "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" - by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) - -subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ - -lemma winding_number_constant: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" - shows "winding_number \ constant_on S" -proof - - have *: "1 \ cmod (winding_number \ y - winding_number \ z)" - if ne: "winding_number \ y \ winding_number \ z" and "y \ S" "z \ S" for y z - proof - - have "winding_number \ y \ \" "winding_number \ z \ \" - using that integer_winding_number [OF \ loop] sg \y \ S\ by auto - with ne show ?thesis - by (auto simp: Ints_def simp flip: of_int_diff) - qed - have cont: "continuous_on S (\w. winding_number \ w)" - using continuous_on_winding_number [OF \] sg - by (meson continuous_on_subset disjoint_eq_subset_Compl) - show ?thesis - using "*" zero_less_one - by (blast intro: continuous_discrete_range_constant [OF cs cont]) -qed - -lemma winding_number_eq: - "\path \; pathfinish \ = pathstart \; w \ S; z \ S; connected S; S \ path_image \ = {}\ - \ winding_number \ w = winding_number \ z" - using winding_number_constant by (metis constant_on_def) - -lemma open_winding_number_levelsets: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - shows "open {z. z \ path_image \ \ winding_number \ z = k}" -proof - - have opn: "open (- path_image \)" - by (simp add: closed_path_image \ open_Compl) - { fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" - obtain e where e: "e>0" "ball z e \ - path_image \" - using open_contains_ball [of "- path_image \"] opn z - by blast - have "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" - apply (rule_tac x=e in exI) - using e apply (simp add: dist_norm ball_def norm_minus_commute) - apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where S = "ball z e"]) - done - } then - show ?thesis - by (auto simp: open_dist) -qed - -subsection\Winding number is zero "outside" a curve\ - -proposition winding_number_zero_in_outside: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" - shows "winding_number \ z = 0" -proof - - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" - using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto - obtain w::complex where w: "w \ ball 0 (B + 1)" - by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) - have "- ball 0 (B + 1) \ outside (path_image \)" - apply (rule outside_subset_convex) - using B subset_ball by auto - then have wout: "w \ outside (path_image \)" - using w by blast - moreover have "winding_number \ constant_on outside (path_image \)" - using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside - by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) - ultimately have "winding_number \ z = winding_number \ w" - by (metis (no_types, hide_lams) constant_on_def z) - also have "\ = 0" - proof - - have wnot: "w \ path_image \" using wout by (simp add: outside_def) - { fix e::real assume "0" "pathfinish p = pathfinish \" - and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" - and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" - using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force - have pip: "path_image p \ ball 0 (B + 1)" - using B - apply (clarsimp simp add: path_image_def dist_norm ball_def) - apply (frule (1) pg1) - apply (fastforce dest: norm_add_less) - done - then have "w \ path_image p" using w by blast - then have "\p. valid_path p \ w \ path_image p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ - (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" - apply (rule_tac x=p in exI) - apply (simp add: p valid_path_polynomial_function) - apply (intro conjI) - using pge apply (simp add: norm_minus_commute) - apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) - apply (rule holomorphic_intros | simp add: dist_norm)+ - using mem_ball_0 w apply blast - using p apply (simp_all add: valid_path_polynomial_function loop pip) - done - } - then show ?thesis - by (auto intro: winding_number_unique [OF \] simp add: winding_number_prop_def wnot) - qed - finally show ?thesis . -qed - -corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" - by (rule winding_number_zero_in_outside) - (auto simp: pathfinish_def pathstart_def path_polynomial_function) - -corollary\<^marker>\tag unimportant\ winding_number_zero_outside: - "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" - by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) - -lemma winding_number_zero_at_infinity: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - shows "\B. \z. B \ norm z \ winding_number \ z = 0" -proof - - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" - using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto - then show ?thesis - apply (rule_tac x="B+1" in exI, clarify) - apply (rule winding_number_zero_outside [OF \ convex_cball [of 0 B] loop]) - apply (meson less_add_one mem_cball_0 not_le order_trans) - using ball_subset_cball by blast -qed - -lemma winding_number_zero_point: - "\path \; convex s; pathfinish \ = pathstart \; open s; path_image \ \ s\ - \ \z. z \ s \ winding_number \ z = 0" - using outside_compact_in_open [of "path_image \" s] path_image_nonempty winding_number_zero_in_outside - by (fastforce simp add: compact_path_image) - - -text\If a path winds round a set, it winds rounds its inside.\ -lemma winding_number_around_inside: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - and cls: "closed s" and cos: "connected s" and s_disj: "s \ path_image \ = {}" - and z: "z \ s" and wn_nz: "winding_number \ z \ 0" and w: "w \ s \ inside s" - shows "winding_number \ w = winding_number \ z" -proof - - have ssb: "s \ inside(path_image \)" - proof - fix x :: complex - assume "x \ s" - hence "x \ path_image \" - by (meson disjoint_iff_not_equal s_disj) - thus "x \ inside (path_image \)" - using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) -qed - show ?thesis - apply (rule winding_number_eq [OF \ loop w]) - using z apply blast - apply (simp add: cls connected_with_inside cos) - apply (simp add: Int_Un_distrib2 s_disj, safe) - by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \ inside_no_overlap) - qed - - -text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ -lemma winding_number_subpath_continuous: - assumes \: "valid_path \" and z: "z \ path_image \" - shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" -proof - - have *: "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = - winding_number (subpath 0 x \) z" - if x: "0 \ x" "x \ 1" for x - proof - - have "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = - 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" - using assms x - apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) - done - also have "\ = winding_number (subpath 0 x \) z" - apply (subst winding_number_valid_path) - using assms x - apply (simp_all add: path_image_subpath valid_path_subpath) - by (force simp: path_image_def) - finally show ?thesis . - qed - show ?thesis - apply (rule continuous_on_eq - [where f = "\x. 1 / (2*pi*\) * - integral {0..x} (\t. 1/(\ t - z) * vector_derivative \ (at t))"]) - apply (rule continuous_intros)+ - apply (rule indefinite_integral_continuous_1) - apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) - using assms - apply (simp add: *) - done -qed - -lemma winding_number_ivt_pos: - assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" - shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" - apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) - apply (rule winding_number_subpath_continuous [OF \ z]) - using assms - apply (auto simp: path_image_def image_def) - done - -lemma winding_number_ivt_neg: - assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" - shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" - apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) - apply (rule winding_number_subpath_continuous [OF \ z]) - using assms - apply (auto simp: path_image_def image_def) - done - -lemma winding_number_ivt_abs: - assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" - shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" - using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] - by force - -lemma winding_number_lt_half_lemma: - assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" - shows "Re(winding_number \ z) < 1/2" -proof - - { assume "Re(winding_number \ z) \ 1/2" - then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" - using winding_number_ivt_pos [OF \ z, of "1/2"] by auto - have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" - using winding_number_exp_2pi [of "subpath 0 t \" z] - apply (simp add: t \ valid_path_imp_path) - using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) - have "b < a \ \ 0" - proof - - have "\ 0 \ {c. b < a \ c}" - by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) - thus ?thesis - by blast - qed - moreover have "b < a \ \ t" - proof - - have "\ t \ {c. b < a \ c}" - by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) - thus ?thesis - by blast - qed - ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az - by (simp add: inner_diff_right)+ - then have False - by (simp add: gt inner_mult_right mult_less_0_iff) - } - then show ?thesis by force -qed - -lemma winding_number_lt_half: - assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" - shows "\Re (winding_number \ z)\ < 1/2" -proof - - have "z \ path_image \" using assms by auto - with assms show ?thesis - apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) - apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] - winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) - done -qed - -lemma winding_number_le_half: - assumes \: "valid_path \" and z: "z \ path_image \" - and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" - shows "\Re (winding_number \ z)\ \ 1/2" -proof - - { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" - have "isCont (winding_number \) z" - by (metis continuous_at_winding_number valid_path_imp_path \ z) - then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" - using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast - define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" - have *: "a \ z' \ b - d / 3 * cmod a" - unfolding z'_def inner_mult_right' divide_inverse - apply (simp add: field_split_simps algebra_simps dot_square_norm power2_eq_square anz) - apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) - done - have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" - using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) - then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" - by simp - then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" - using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp - then have wnz_12': "\Re (winding_number \ z')\ > 1/2" - by linarith - moreover have "\Re (winding_number \ z')\ < 1/2" - apply (rule winding_number_lt_half [OF \ *]) - using azb \d>0\ pag - apply (auto simp: add_strict_increasing anz field_split_simps dest!: subsetD) - done - ultimately have False - by simp - } - then show ?thesis by force -qed - -lemma winding_number_lt_half_linepath: "z \ closed_segment a b \ \Re (winding_number (linepath a b) z)\ < 1/2" - using separating_hyperplane_closed_point [of "closed_segment a b" z] - apply auto - apply (simp add: closed_segment_def) - apply (drule less_imp_le) - apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) - apply (auto simp: segment) - done - - -text\ Positivity of WN for a linepath.\ -lemma winding_number_linepath_pos_lt: - assumes "0 < Im ((b - a) * cnj (b - z))" - shows "0 < Re(winding_number(linepath a b) z)" -proof - - have z: "z \ path_image (linepath a b)" - using assms - by (simp add: closed_segment_def) (force simp: algebra_simps) - show ?thesis - apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) - apply (simp add: linepath_def algebra_simps) - done -qed - - -subsection\Cauchy's integral formula, again for a convex enclosing set\ - -lemma Cauchy_integral_formula_weak: - assumes s: "convex s" and "finite k" and conf: "continuous_on s f" - and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" - and z: "z \ interior s - k" and vpg: "valid_path \" - and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - obtain f' where f': "(f has_field_derivative f') (at z)" - using fcd [OF z] by (auto simp: field_differentiable_def) - have pas: "path_image \ \ s" and znotin: "z \ path_image \" using pasz by blast+ - have c: "continuous (at x within s) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ s" for x - proof (cases "x = z") - case True then show ?thesis - 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 - next - case False - then have dxz: "dist x z > 0" by auto - have cf: "continuous (at x within s) f" - using conf continuous_on_eq_continuous_within that by blast - have "continuous (at x within s) (\w. (f w - f z) / (w - z))" - by (rule cf continuous_intros | simp add: False)+ - then show ?thesis - apply (rule continuous_transform_within [OF _ dxz that, of "\w::complex. (f w - f z)/(w - z)"]) - 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 - show ?thesis - apply (rule has_contour_integral_eq) - using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] - apply (auto simp: ac_simps divide_simps) - done -qed - -theorem Cauchy_integral_formula_convex_simple: - "\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) - subsection\Homotopy forms of Cauchy's theorem\ lemma Cauchy_theorem_homotopic: assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" proof - have pathsf: "linked_paths atends g h" using hom by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop) obtain k :: "real \ real \ complex" where contk: "continuous_on ({0..1} \ {0..1}) k" and ks: "k ` ({0..1} \ {0..1}) \ s" and k [simp]: "\x. k (0, x) = g x" "\x. k (1, x) = h x" and ksf: "\t\{0..1}. linked_paths atends g (\x. k (t, x))" using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm) have ucontk: "uniformly_continuous_on ({0..1} \ {0..1}) k" by (blast intro: compact_Times compact_uniformly_continuous [OF contk]) { fix t::real assume t: "t \ {0..1}" have pak: "path (k \ (\u. (t, u)))" unfolding path_def apply (rule continuous_intros continuous_on_subset [OF contk])+ using t by force have pik: "path_image (k \ Pair t) \ s" using ks t by (auto simp: path_image_def) obtain e where "e>0" and e: "\g h. \valid_path g; valid_path h; \u\{0..1}. cmod (g u - (k \ Pair t) u) < e \ cmod (h u - (k \ Pair t) u) < e; linked_paths atends g h\ \ contour_integral h f = contour_integral g f" using contour_integral_nearby [OF \open s\ pak pik, of atends] f by metis obtain d where "d>0" and d: "\x x'. \x \ {0..1} \ {0..1}; x' \ {0..1} \ {0..1}; norm (x'-x) < d\ \ norm (k x' - k x) < e/4" by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \e>0\) { fix t1 t2 assume t1: "0 \ t1" "t1 \ 1" and t2: "0 \ t2" "t2 \ 1" and ltd: "\t1 - t\ < d" "\t2 - t\ < d" have no2: "\g1 k1 kt. \norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\ \ norm(g1 - kt) < e" using \e > 0\ apply (rule_tac y = k1 in norm_triangle_half_l) apply (auto simp: norm_minus_commute intro: order_less_trans) done have "\d>0. \g1 g2. valid_path g1 \ valid_path g2 \ (\u\{0..1}. cmod (g1 u - k (t1, u)) < d \ cmod (g2 u - k (t2, u)) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f" apply (rule_tac x="e/4" in exI) using t t1 t2 ltd \e > 0\ apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1) done } then have "\e. 0 < e \ (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < e \ \t2 - t\ < e \ (\d. 0 < d \ (\g1 g2. valid_path g1 \ valid_path g2 \ (\u \ {0..1}. norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f)))" by (rule_tac x=d in exI) (simp add: \d > 0\) } then obtain ee where ee: "\t. t \ {0..1} \ ee t > 0 \ (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < ee t \ \t2 - t\ < ee t \ (\d. 0 < d \ (\g1 g2. valid_path g1 \ valid_path g2 \ (\u \ {0..1}. norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f)))" by metis note ee_rule = ee [THEN conjunct2, rule_format] define C where "C = (\t. ball t (ee t / 3)) ` {0..1}" obtain C' where C': "C' \ C" "finite C'" and C'01: "{0..1} \ \C'" proof (rule compactE [OF compact_interval]) show "{0..1} \ \C" using ee [THEN conjunct1] by (auto simp: C_def dist_norm) qed (use C_def in auto) define kk where "kk = {t \ {0..1}. ball t (ee t / 3) \ C'}" have kk01: "kk \ {0..1}" by (auto simp: kk_def) define e where "e = Min (ee ` kk)" have C'_eq: "C' = (\t. ball t (ee t / 3)) ` kk" using C' by (auto simp: kk_def C_def) have ee_pos[simp]: "\t. t \ {0..1} \ ee t > 0" by (simp add: kk_def ee) moreover have "finite kk" using \finite C'\ kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD) moreover have "kk \ {}" using \{0..1} \ \C'\ C'_eq by force ultimately have "e > 0" using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def) then obtain N::nat where "N > 0" and N: "1/N < e/3" by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral) have e_le_ee: "\i. i \ kk \ e \ ee i" using \finite kk\ by (simp add: e_def Min_le_iff [of "ee ` kk"]) have plus: "\t \ kk. x \ ball t (ee t / 3)" if "x \ {0..1}" for x using C' subsetD [OF C'01 that] unfolding C'_eq by blast have [OF order_refl]: "\d. 0 < d \ (\j. valid_path j \ (\u \ {0..1}. norm(j u - k (n/N, u)) < d) \ linked_paths atends g j \ contour_integral j f = contour_integral g f)" if "n \ N" for n using that proof (induct n) case 0 show ?case using ee_rule [of 0 0 0] apply clarsimp apply (rule_tac x=d in exI, safe) by (metis diff_self vpg norm_zero) next case (Suc n) then have N01: "n/N \ {0..1}" "(Suc n)/N \ {0..1}" by auto then obtain t where t: "t \ kk" "n/N \ ball t (ee t / 3)" using plus [of "n/N"] by blast then have nN_less: "\n/N - t\ < ee t" by (simp add: dist_norm del: less_divide_eq_numeral1) have n'N_less: "\real (Suc n) / real N - t\ < ee t" using t N \N > 0\ e_le_ee [of t] by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) have t01: "t \ {0..1}" using \kk \ {0..1}\ \t \ kk\ by blast obtain d1 where "d1 > 0" and d1: "\g1 g2. \valid_path g1; valid_path g2; \u\{0..1}. cmod (g1 u - k (n/N, u)) < d1 \ cmod (g2 u - k ((Suc n) / N, u)) < d1; linked_paths atends g1 g2\ \ contour_integral g2 f = contour_integral g1 f" using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce have "n \ N" using Suc.prems by auto with Suc.hyps obtain d2 where "d2 > 0" and d2: "\j. \valid_path j; \u\{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\ \ contour_integral j f = contour_integral g f" by auto have "continuous_on {0..1} (k \ (\u. (n/N, u)))" apply (rule continuous_intros continuous_on_subset [OF contk])+ using N01 by auto then have pkn: "path (\u. k (n/N, u))" by (simp add: path_def) have min12: "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) obtain p where "polynomial_function p" and psf: "pathstart p = pathstart (\u. k (n/N, u))" "pathfinish p = pathfinish (\u. k (n/N, u))" and pk_le: "\t. t\{0..1} \ cmod (p t - k (n/N, t)) < min d1 d2" using path_approx_polynomial_function [OF pkn min12] by blast then have vpp: "valid_path p" using valid_path_polynomial_function by blast have lpa: "linked_paths atends g p" by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) show ?case proof (intro exI; safe) fix j assume "valid_path j" "linked_paths atends g j" and "\u\{0..1}. cmod (j u - k (real (Suc n) / real N, u)) < min d1 d2" then have "contour_integral j f = contour_integral p f" using pk_le N01(1) ksf by (force intro!: vpp d1 simp add: linked_paths_def psf) also have "... = contour_integral g f" using pk_le by (force intro!: vpp d2 lpa) finally show "contour_integral j f = contour_integral g f" . qed (simp add: \0 < d1\ \0 < d2\) qed then obtain d where "0 < d" "\j. valid_path j \ (\u \ {0..1}. norm(j u - k (1,u)) < d) \ linked_paths atends g j \ contour_integral j f = contour_integral g f" using \N>0\ by auto then have "linked_paths atends g h \ contour_integral h f = contour_integral g f" using \N>0\ vph by fastforce then show ?thesis by (simp add: pathsf) qed proposition Cauchy_theorem_homotopic_paths: assumes hom: "homotopic_paths s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" using Cauchy_theorem_homotopic [of True s g h] assms by simp proposition Cauchy_theorem_homotopic_loops: assumes hom: "homotopic_loops s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" using Cauchy_theorem_homotopic [of False s g h] assms by simp lemma has_contour_integral_newpath: "\(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\ \ (f has_contour_integral y) g" using has_contour_integral_integral contour_integral_unique by auto lemma Cauchy_theorem_null_homotopic: "\f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\ \ (f has_contour_integral 0) g" apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp) using contour_integrable_holomorphic_simple apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset) by (simp add: Cauchy_theorem_homotopic_loops) -subsection\<^marker>\tag unimportant\ \More winding number properties\ - -text\including the fact that it's +-1 inside a simple closed curve.\ - -lemma winding_number_homotopic_paths: - assumes "homotopic_paths (-{z}) g h" - shows "winding_number g z = winding_number h z" -proof - - have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto - moreover have pag: "z \ path_image g" and pah: "z \ path_image h" - using homotopic_paths_imp_subset [OF assms] by auto - ultimately obtain d e where "d > 0" "e > 0" - and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ - \ homotopic_paths (-{z}) g p" - and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ - \ homotopic_paths (-{z}) h q" - using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force - obtain p where p: - "valid_path p" "z \ path_image p" - "pathstart p = pathstart g" "pathfinish p = pathfinish g" - and gp_less:"\t\{0..1}. cmod (g t - p t) < d" - and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" - using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast - obtain q where q: - "valid_path q" "z \ path_image q" - "pathstart q = pathstart h" "pathfinish q = pathfinish h" - and hq_less: "\t\{0..1}. cmod (h t - q t) < e" - and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" - using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast - have "homotopic_paths (- {z}) g p" - by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) - moreover have "homotopic_paths (- {z}) h q" - by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) - ultimately have "homotopic_paths (- {z}) p q" - by (blast intro: homotopic_paths_trans homotopic_paths_sym assms) - then have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" - by (rule Cauchy_theorem_homotopic_paths) (auto intro!: holomorphic_intros simp: p q) - then show ?thesis - by (simp add: pap paq) -qed - -lemma winding_number_homotopic_loops: - assumes "homotopic_loops (-{z}) g h" - shows "winding_number g z = winding_number h z" -proof - - have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto - moreover have pag: "z \ path_image g" and pah: "z \ path_image h" - using homotopic_loops_imp_subset [OF assms] by auto - moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" - using homotopic_loops_imp_loop [OF assms] by auto - ultimately obtain d e where "d > 0" "e > 0" - and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ - \ homotopic_loops (-{z}) g p" - and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ - \ homotopic_loops (-{z}) h q" - using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force - obtain p where p: - "valid_path p" "z \ path_image p" - "pathstart p = pathstart g" "pathfinish p = pathfinish g" - and gp_less:"\t\{0..1}. cmod (g t - p t) < d" - and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" - using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast - obtain q where q: - "valid_path q" "z \ path_image q" - "pathstart q = pathstart h" "pathfinish q = pathfinish h" - and hq_less: "\t\{0..1}. cmod (h t - q t) < e" - and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" - using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast - have gp: "homotopic_loops (- {z}) g p" - by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) - have hq: "homotopic_loops (- {z}) h q" - by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) - have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" - proof (rule Cauchy_theorem_homotopic_loops) - show "homotopic_loops (- {z}) p q" - by (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) - qed (auto intro!: holomorphic_intros simp: p q) - then show ?thesis - by (simp add: pap paq) -qed - -lemma winding_number_paths_linear_eq: - "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; - \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ - \ winding_number h z = winding_number g z" - by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths) - -lemma winding_number_loops_linear_eq: - "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; - \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ - \ winding_number h z = winding_number g z" - by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops) - -lemma winding_number_nearby_paths_eq: - "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; - \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ - \ winding_number h z = winding_number g z" - by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) - -lemma winding_number_nearby_loops_eq: - "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; - \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ - \ winding_number h z = winding_number g z" - by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) - - -lemma winding_number_subpath_combine: - "\path g; z \ path_image g; - u \ {0..1}; v \ {0..1}; w \ {0..1}\ - \ winding_number (subpath u v g) z + winding_number (subpath v w g) z = - winding_number (subpath u w g) z" -apply (rule trans [OF winding_number_join [THEN sym] - winding_number_homotopic_paths [OF homotopic_join_subpaths]]) - using path_image_subpath_subset by auto - -subsection\Partial circle path\ - -definition\<^marker>\tag important\ part_circlepath :: "[complex, real, real, real, real] \ complex" - where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" - -lemma pathstart_part_circlepath [simp]: - "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" -by (metis part_circlepath_def pathstart_def pathstart_linepath) - -lemma pathfinish_part_circlepath [simp]: - "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" -by (metis part_circlepath_def pathfinish_def pathfinish_linepath) - -lemma reversepath_part_circlepath[simp]: - "reversepath (part_circlepath z r s t) = part_circlepath z r t s" - unfolding part_circlepath_def reversepath_def linepath_def - by (auto simp:algebra_simps) - -lemma has_vector_derivative_part_circlepath [derivative_intros]: - "((part_circlepath z r s t) has_vector_derivative - (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) - (at x within X)" - apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) - apply (rule has_vector_derivative_real_field) - apply (rule derivative_eq_intros | simp)+ - done - -lemma differentiable_part_circlepath: - "part_circlepath c r a b differentiable at x within A" - using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast - -lemma vector_derivative_part_circlepath: - "vector_derivative (part_circlepath z r s t) (at x) = - \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" - using has_vector_derivative_part_circlepath vector_derivative_at by blast - -lemma vector_derivative_part_circlepath01: - "\0 \ x; x \ 1\ - \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = - \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" - using has_vector_derivative_part_circlepath - by (auto simp: vector_derivative_at_within_ivl) - -lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" - apply (simp add: valid_path_def) - apply (rule C1_differentiable_imp_piecewise) - apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath - intro!: continuous_intros) - done - -lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" - by (simp add: valid_path_imp_path) - -proposition path_image_part_circlepath: - assumes "s \ t" - shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" -proof - - { fix z::real - assume "0 \ z" "z \ 1" - with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" - apply (rule_tac x="(1 - z) * s + z * t" in exI) - apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) - apply (rule conjI) - using mult_right_mono apply blast - using affine_ineq by (metis "mult.commute") - } - moreover - { fix z - assume "s \ z" "z \ t" - then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" - apply (rule_tac x="(z - s)/(t - s)" in image_eqI) - apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) - apply (auto simp: field_split_simps) - done - } - ultimately show ?thesis - by (fastforce simp add: path_image_def part_circlepath_def) -qed - -lemma path_image_part_circlepath': - "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" -proof - - have "path_image (part_circlepath z r s t) = - (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" - by (simp add: image_image path_image_def part_circlepath_def) - also have "linepath s t ` {0..1} = closed_segment s t" - by (rule linepath_image_01) - finally show ?thesis by (simp add: cis_conv_exp) -qed - -lemma path_image_part_circlepath_subset: - "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" -by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) - -lemma in_path_image_part_circlepath: - assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" - shows "norm(w - z) = r" -proof - - have "w \ {c. dist z c = r}" - by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) - thus ?thesis - by (simp add: dist_norm norm_minus_commute) -qed - -lemma path_image_part_circlepath_subset': - assumes "r \ 0" - shows "path_image (part_circlepath z r s t) \ sphere z r" -proof (cases "s \ t") - case True - thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp -next - case False - thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms - by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all -qed - -lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" - by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) - -lemma contour_integral_bound_part_circlepath: - assumes "f contour_integrable_on part_circlepath c r a b" - assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" - shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" -proof - - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * - exp (\ * linepath a b x))" - have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" - proof (rule integral_norm_bound_integral, goal_cases) - case 1 - with assms(1) show ?case - by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) - next - case (3 x) - with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult - by (intro mult_mono) (auto simp: path_image_def) - qed auto - also have "?I = contour_integral (part_circlepath c r a b) f" - by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) - finally show ?thesis by simp -qed - -lemma has_contour_integral_part_circlepath_iff: - assumes "a < b" - shows "(f has_contour_integral I) (part_circlepath c r a b) \ - ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" -proof - - have "(f has_contour_integral I) (part_circlepath c r a b) \ - ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) - (at x within {0..1})) has_integral I) {0..1}" - unfolding has_contour_integral_def .. - also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * - cis (linepath a b x)) has_integral I) {0..1}" - by (intro has_integral_cong, subst vector_derivative_part_circlepath01) - (simp_all add: cis_conv_exp) - also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * - r * \ * exp (\ * linepath (of_real a) (of_real b) x) * - vector_derivative (linepath (of_real a) (of_real b)) - (at x within {0..1})) has_integral I) {0..1}" - by (intro has_integral_cong, subst vector_derivative_linepath_within) - (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) - also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) - (linepath (of_real a) (of_real b))" - by (simp add: has_contour_integral_def) - also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms - by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) - finally show ?thesis . -qed - -lemma contour_integrable_part_circlepath_iff: - assumes "a < b" - shows "f contour_integrable_on (part_circlepath c r a b) \ - (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" - using assms by (auto simp: contour_integrable_on_def integrable_on_def - has_contour_integral_part_circlepath_iff) - -lemma contour_integral_part_circlepath_eq: - assumes "a < b" - shows "contour_integral (part_circlepath c r a b) f = - integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" -proof (cases "f contour_integrable_on part_circlepath c r a b") - case True - hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" - using assms by (simp add: contour_integrable_part_circlepath_iff) - with True show ?thesis - using has_contour_integral_part_circlepath_iff[OF assms] - contour_integral_unique has_integral_integrable_integral by blast -next - case False - hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" - using assms by (simp add: contour_integrable_part_circlepath_iff) - with False show ?thesis - by (simp add: not_integrable_contour_integral not_integrable_integral) -qed - -lemma contour_integral_part_circlepath_reverse: - "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" - by (subst reversepath_part_circlepath [symmetric], subst contour_integral_reversepath) simp_all - -lemma contour_integral_part_circlepath_reverse': - "b < a \ contour_integral (part_circlepath c r a b) f = - -contour_integral (part_circlepath c r b a) f" - by (rule contour_integral_part_circlepath_reverse) - -lemma finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" -proof (cases "w = 0") - case True then show ?thesis by auto -next - case False - have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" - apply (simp add: norm_mult finite_int_iff_bounded_le) - apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) - apply (auto simp: field_split_simps le_floor_iff) - done - have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" - by blast - show ?thesis - apply (subst exp_Ln [OF False, symmetric]) - apply (simp add: exp_eq) - using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) - done -qed - -lemma finite_bounded_log2: - fixes a::complex - assumes "a \ 0" - shows "finite {z. norm z \ b \ exp(a*z) = w}" -proof - - have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" - by (rule finite_imageI [OF finite_bounded_log]) - show ?thesis - by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) -qed - -lemma has_contour_integral_bound_part_circlepath_strong: - assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" - and "finite k" and le: "0 \ B" "0 < r" "s \ t" - and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" - shows "cmod i \ B * r * (t - s)" -proof - - consider "s = t" | "s < t" using \s \ t\ by linarith - then show ?thesis - proof cases - case 1 with fi [unfolded has_contour_integral] - have "i = 0" by (simp add: vector_derivative_part_circlepath) - with assms show ?thesis by simp - next - case 2 - have [simp]: "\r\ = r" using \r > 0\ by linarith - have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" - by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) - have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y - proof - - define w where "w = (y - z)/of_real r / exp(\ * of_real s)" - have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" - apply (rule finite_vimageI [OF finite_bounded_log2]) - using \s < t\ apply (auto simp: inj_of_real) - done - show ?thesis - apply (simp add: part_circlepath_def linepath_def vimage_def) - apply (rule finite_subset [OF _ fin]) - using le - apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) - done - qed - then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" - by (rule finite_finite_vimage_IntI [OF \finite k\]) - have **: "((\x. if (part_circlepath z r s t x) \ k then 0 - else f(part_circlepath z r s t x) * - vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" - by (rule has_integral_spike [OF negligible_finite [OF fin01]]) (use fi has_contour_integral in auto) - have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" - by (auto intro!: B [unfolded path_image_def image_def, simplified]) - show ?thesis - apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) - using assms apply force - apply (simp add: norm_mult vector_derivative_part_circlepath) - using le * "2" \r > 0\ by auto - qed -qed - -lemma has_contour_integral_bound_part_circlepath: - "\(f has_contour_integral i) (part_circlepath z r s t); - 0 \ B; 0 < r; s \ t; - \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ - \ norm i \ B*r*(t - s)" - by (auto intro: has_contour_integral_bound_part_circlepath_strong) - -lemma contour_integrable_continuous_part_circlepath: - "continuous_on (path_image (part_circlepath z r s t)) f - \ f contour_integrable_on (part_circlepath z r s t)" - apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) - apply (rule integrable_continuous_real) - apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) - done - -proposition winding_number_part_circlepath_pos_less: - assumes "s < t" and no: "norm(w - z) < r" - shows "0 < Re (winding_number(part_circlepath z r s t) w)" -proof - - have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) - note valid_path_part_circlepath - moreover have " w \ path_image (part_circlepath z r s t)" - using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) - moreover have "0 < r * (t - s) * (r - cmod (w - z))" - using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) - ultimately show ?thesis - apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) - apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) - apply (rule mult_left_mono)+ - using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] - apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) - using assms \0 < r\ by auto -qed - -lemma simple_path_part_circlepath: - "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" -proof (cases "r = 0 \ s = t") - case True - then show ?thesis - unfolding part_circlepath_def simple_path_def - by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ -next - case False then have "r \ 0" "s \ t" by auto - have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" - by (simp add: algebra_simps) - have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 - \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" - by auto - have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ - (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" - by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] - intro: exI [where x = "-n" for n]) - have 1: "\s - t\ \ 2 * pi" - if "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1" - proof (rule ccontr) - assume "\ \s - t\ \ 2 * pi" - then have *: "\n. t - s \ of_int n * \s - t\" - using False that [of "2*pi / \t - s\"] - by (simp add: abs_minus_commute divide_simps) - show False - using * [of 1] * [of "-1"] by auto - qed - have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n - proof - - have "t-s = 2 * (real_of_int n * pi)/x" - using that by (simp add: field_simps) - then show ?thesis by (metis abs_minus_commute) - qed - have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" - by force - show ?thesis using False - apply (simp add: simple_path_def) - apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) - apply (subst abs_away) - apply (auto simp: 1) - apply (rule ccontr) - apply (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) - done -qed - -lemma arc_part_circlepath: - assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" - shows "arc (part_circlepath z r s t)" -proof - - have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" - and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n - proof (rule ccontr) - assume "x \ y" - have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" - by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_div_cancel_left eq) - then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" - by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) - with \x \ y\ have st: "s-t = (of_int n * (pi * 2) / (y-x))" - by (force simp: field_simps) - have "\real_of_int n\ < \y - x\" - using assms \x \ y\ by (simp add: st abs_mult field_simps) - then show False - using assms x y st by (auto dest: of_int_lessD) - qed - show ?thesis - using assms - apply (simp add: arc_def) - apply (simp add: part_circlepath_def inj_on_def exp_eq) - apply (blast intro: *) - done -qed - -subsection\Special case of one complete circle\ - -definition\<^marker>\tag important\ circlepath :: "[complex, real, real] \ complex" - where "circlepath z r \ part_circlepath z r 0 (2*pi)" - -lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" - by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) - -lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" - by (simp add: circlepath_def) - -lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" - by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) - -lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" -proof - - have "z + of_real r * exp (2 * pi * \ * (x + 1/2)) = - z + of_real r * exp (2 * pi * \ * x + pi * \)" - by (simp add: divide_simps) (simp add: algebra_simps) - also have "\ = z - r * exp (2 * pi * \ * x)" - by (simp add: exp_add) - finally show ?thesis - by (simp add: circlepath path_image_def sphere_def dist_norm) -qed - -lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" - using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] - by (simp add: add.commute) - -lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" - using circlepath_add1 [of z r "x-1/2"] - by (simp add: add.commute) - -lemma path_image_circlepath_minus_subset: - "path_image (circlepath z (-r)) \ path_image (circlepath z r)" - apply (simp add: path_image_def image_def circlepath_minus, clarify) - apply (case_tac "xa \ 1/2", force) - apply (force simp: circlepath_add_half)+ - done - -lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" - using path_image_circlepath_minus_subset by fastforce - -lemma has_vector_derivative_circlepath [derivative_intros]: - "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * of_real x))) - (at x within X)" - apply (simp add: circlepath_def scaleR_conv_of_real) - apply (rule derivative_eq_intros) - apply (simp add: algebra_simps) - done - -lemma vector_derivative_circlepath: - "vector_derivative (circlepath z r) (at x) = - 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" -using has_vector_derivative_circlepath vector_derivative_at by blast - -lemma vector_derivative_circlepath01: - "\0 \ x; x \ 1\ - \ vector_derivative (circlepath z r) (at x within {0..1}) = - 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" - using has_vector_derivative_circlepath - by (auto simp: vector_derivative_at_within_ivl) - -lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" - by (simp add: circlepath_def) - -lemma path_circlepath [simp]: "path (circlepath z r)" - by (simp add: valid_path_imp_path) - -lemma path_image_circlepath_nonneg: - assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" -proof - - have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x - proof (cases "x = z") - case True then show ?thesis by force - next - case False - define w where "w = x - z" - then have "w \ 0" by (simp add: False) - have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" - using cis_conv_exp complex_eq_iff by auto - show ?thesis - apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) - apply (simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) - apply (rule_tac x="t / (2*pi)" in image_eqI) - apply (simp add: field_simps \w \ 0\) - using False ** - apply (auto simp: w_def) - done - qed - show ?thesis - unfolding circlepath path_image_def sphere_def dist_norm - by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) -qed - -lemma path_image_circlepath [simp]: - "path_image (circlepath z r) = sphere z \r\" - using path_image_circlepath_minus - by (force simp: path_image_circlepath_nonneg abs_if) - -lemma has_contour_integral_bound_circlepath_strong: - "\(f has_contour_integral i) (circlepath z r); - finite k; 0 \ B; 0 < r; - \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ - \ norm i \ B*(2*pi*r)" - unfolding circlepath_def - by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) - -lemma has_contour_integral_bound_circlepath: - "\(f has_contour_integral i) (circlepath z r); - 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ - \ norm i \ B*(2*pi*r)" - by (auto intro: has_contour_integral_bound_circlepath_strong) - -lemma contour_integrable_continuous_circlepath: - "continuous_on (path_image (circlepath z r)) f - \ f contour_integrable_on (circlepath z r)" - by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) - -lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" - by (simp add: circlepath_def simple_path_part_circlepath) - -lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" - by (simp add: sphere_def dist_norm norm_minus_commute) - -lemma contour_integral_circlepath: - assumes "r > 0" - shows "contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" -proof (rule contour_integral_unique) - show "((\w. 1 / (w - z)) has_contour_integral 2 * complex_of_real pi * \) (circlepath z r)" - unfolding has_contour_integral_def using assms - apply (subst has_integral_cong) - apply (simp add: vector_derivative_circlepath01) - using has_integral_const_real [of _ 0 1] apply (force simp: circlepath) - done -qed - -lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" - apply (rule winding_number_unique_loop) - apply (simp_all add: sphere_def valid_path_imp_path) - apply (rule_tac x="circlepath z r" in exI) - apply (simp add: sphere_def contour_integral_circlepath) - done - -proposition winding_number_circlepath: - assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" -proof (cases "w = z") - case True then show ?thesis - using assms winding_number_circlepath_centre by auto -next - case False - have [simp]: "r > 0" - using assms le_less_trans norm_ge_zero by blast - define r' where "r' = norm(w - z)" - have "r' < r" - by (simp add: assms r'_def) - have disjo: "cball z r' \ sphere z r = {}" - using \r' < r\ by (force simp: cball_def sphere_def) - have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" - proof (rule winding_number_around_inside [where s = "cball z r'"]) - show "winding_number (circlepath z r) z \ 0" - by (simp add: winding_number_circlepath_centre) - show "cball z r' \ path_image (circlepath z r) = {}" - by (simp add: disjo less_eq_real_def) - qed (auto simp: r'_def dist_norm norm_minus_commute) - also have "\ = 1" - by (simp add: winding_number_circlepath_centre) - finally show ?thesis . -qed - - -text\ Hence the Cauchy formula for points inside a circle.\ - -theorem Cauchy_integral_circlepath: - assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r" - shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) - (circlepath z r)" -proof - - have "r > 0" - using assms le_less_trans norm_ge_zero by blast - have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) - (circlepath z r)" - proof (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) - show "\x. x \ interior (cball z r) - {} \ - f field_differentiable at x" - using holf holomorphic_on_imp_differentiable_at by auto - have "w \ sphere z r" - by simp (metis dist_commute dist_norm not_le order_refl wz) - then show "path_image (circlepath z r) \ cball z r - {w}" - using \r > 0\ by (auto simp add: cball_def sphere_def) - qed (use wz in \simp_all add: dist_norm norm_minus_commute contf\) - then show ?thesis - by (simp add: winding_number_circlepath assms) -qed - -corollary\<^marker>\tag unimportant\ Cauchy_integral_circlepath_simple: - assumes "f holomorphic_on cball z r" "norm(w - z) < r" - shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) - (circlepath z r)" -using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) - - -lemma no_bounded_connected_component_imp_winding_number_zero: - assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" - and nb: "\z. bounded (connected_component_set (- s) z) \ z \ s" - shows "winding_number g z = 0" -apply (rule winding_number_zero_in_outside) -apply (simp_all add: assms) -by (metis nb [of z] \path_image g \ s\ \z \ s\ contra_subsetD mem_Collect_eq outside outside_mono) - -lemma no_bounded_path_component_imp_winding_number_zero: - assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" - and nb: "\z. bounded (path_component_set (- s) z) \ z \ s" - shows "winding_number g z = 0" -apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) -by (simp add: bounded_subset nb path_component_subset_connected_component) - - -subsection\ Uniform convergence of path integral\ - -text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ - -proposition contour_integral_uniform_limit: - assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" - and ul_f: "uniform_limit (path_image \) f l F" - and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" - and \: "valid_path \" - and [simp]: "\ trivial_limit F" - shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" -proof - - have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) - { fix e::real - assume "0 < e" - then have "0 < e / (\B\ + 1)" by simp - then have "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" - using ul_f [unfolded uniform_limit_iff dist_norm] by auto - with ev_fint - obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" - and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" - using eventually_happens [OF eventually_conj] - by (fastforce simp: contour_integrable_on path_image_def) - have Ble: "B * e / (\B\ + 1) \ e" - using \0 \ B\ \0 < e\ by (simp add: field_split_simps) - have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" - proof (intro exI conjI ballI) - show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" - if "x \ {0..1}" for x - apply (rule order_trans [OF _ Ble]) - using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ - apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) - apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le]) - done - qed (rule inta) - } - then show lintg: "l contour_integrable_on \" - unfolding contour_integrable_on by (metis (mono_tags, lifting)integrable_uniform_limit_real) - { fix e::real - define B' where "B' = B + 1" - have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) - assume "0 < e" - then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" - using ul_f [unfolded uniform_limit_iff dist_norm, rule_format, of "e / B' / 2"] B' - by (simp add: field_simps) - have ie: "integral {0..1::real} (\x. e / 2) < e" using \0 < e\ by simp - have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e / 2" - if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t - proof - - have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" - using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto - also have "\ < e" - by (simp add: B' \0 < e\ mult_imp_div_pos_less) - finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . - then show ?thesis - by (simp add: left_diff_distrib [symmetric] norm_mult) - qed - have le_e: "\x. \\xa\{0..1}. 2 * cmod (f x (\ xa) - l (\ xa)) < e / B'; f x contour_integrable_on \\ - \ cmod (integral {0..1} - (\u. f x (\ u) * vector_derivative \ (at u) - l (\ u) * vector_derivative \ (at u))) < e" - apply (rule le_less_trans [OF integral_norm_bound_integral ie]) - apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) - apply (blast intro: *)+ - done - have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" - apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) - apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) - apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric] le_e) - done - } - then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" - by (rule tendstoI) -qed - -corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: - assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" - and "uniform_limit (sphere z r) f l F" - and "\ trivial_limit F" "0 < r" - shows "l contour_integrable_on (circlepath z r)" - "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" - using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) - - -subsection\<^marker>\tag unimportant\ \General stepping result for derivative formulas\ - -lemma Cauchy_next_derivative: - assumes "continuous_on (path_image \) f'" - and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" - and int: "\w. w \ s - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" - and k: "k \ 0" - and "open s" - and \: "valid_path \" - and w: "w \ s - path_image \" - shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" - and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) - (at w)" (is "?thes2") -proof - - have "open (s - path_image \)" using \open s\ closed_valid_path_image \ by blast - then obtain d where "d>0" and d: "ball w d \ s - path_image \" using w - using open_contains_ball by blast - have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" - by (metis norm_of_nat of_nat_Suc) - have cint: "\x. \x \ w; cmod (x - w) < d\ - \ (\z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \" - 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)+ - have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) - contour_integrable_on \" - unfolding eventually_at - apply (rule_tac x=d in exI) - apply (simp add: \d > 0\ dist_norm field_simps cint) - done - have bim_g: "bounded (image f' (path_image \))" - by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) - then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" - by (force simp: bounded_pos path_image_def) - have twom: "\\<^sub>F n in at w. - \x\path_image \. - cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" - if "0 < e" for e - proof - - have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" - if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" - and uw_less: "cmod (u - w) < e * (d/2) ^ (k+2) / (1 + real k)" - for u x - proof - - define ff where [abs_def]: - "ff n w = - (if n = 0 then inverse(x - w)^k - else if n = 1 then k / (x - w)^(Suc k) - else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w - have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" - by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) - have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d/2))" - if "z \ ball w (d/2)" "i \ 1" for i z - proof - - have "z \ path_image \" - using \x \ path_image \\ d that ball_divide_subset_numeral by blast - then have xz[simp]: "x \ z" using \x \ path_image \\ by blast - then have neq: "x * x + z * z \ x * (z * 2)" - by (blast intro: dest!: sum_sqs_eq) - with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto - then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" - by (simp add: algebra_simps) - show ?thesis using \i \ 1\ - apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) - apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ - done - qed - { fix a::real and b::real assume ab: "a > 0" "b > 0" - then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" - by (subst mult_le_cancel_left_pos) - (use \k \ 0\ in \auto simp: divide_simps\) - with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" - by (simp add: field_simps) - } note canc = this - have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d/2) ^ (k + 2)" - if "v \ ball w (d/2)" for v - proof - - have lessd: "\z. cmod (\ z - v) < d/2 \ cmod (w - \ z) < d" - by (metis that norm_minus_commute norm_triangle_half_r dist_norm mem_ball) - have "d/2 \ cmod (x - v)" using d x that - using lessd d x - by (auto simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps) - then have "d \ cmod (x - v) * 2" - by (simp add: field_split_simps) - then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" - using \0 < d\ order_less_imp_le power_mono by blast - have "x \ v" using that - using \x \ path_image \\ ball_divide_subset_numeral d by fastforce - then show ?thesis - using \d > 0\ apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) - using dpow_le apply (simp add: field_split_simps) - done - qed - have ub: "u \ ball w (d/2)" - using uwd by (simp add: dist_commute dist_norm) - have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d/2) ^ k))" - using complex_Taylor [OF _ ff1 ff2 _ ub, of w, simplified] - by (simp add: ff_def \0 < d\) - then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" - by (simp add: field_simps) - then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - / (cmod (u - w) * real k) - \ (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" - using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) - also have "\ < e" - using uw_less \0 < d\ by (simp add: mult_ac divide_simps) - finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) - / cmod ((u - w) * real k) < e" - by (simp add: norm_mult) - have "x \ u" - using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) - show ?thesis - apply (rule le_less_trans [OF _ e]) - using \k \ 0\ \x \ u\ \u \ w\ - apply (simp add: field_simps norm_divide [symmetric]) - done - qed - show ?thesis - unfolding eventually_at - apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) - apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) - done - qed - have 2: "uniform_limit (path_image \) (\n x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) (\x. f' x / (x - w) ^ Suc k) (at w)" - unfolding uniform_limit_iff dist_norm - proof clarify - fix e::real - assume "0 < e" - have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" - if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" - and x: "0 \ x" "x \ 1" - for u x - proof (cases "(f' (\ x)) = 0") - case True then show ?thesis by (simp add: \0 < e\) - next - case False - have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = - cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k))" - by (simp add: field_simps) - also have "\ = cmod (f' (\ x)) * - cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k)" - by (simp add: norm_mult) - also have "\ < cmod (f' (\ x)) * (e/C)" - using False mult_strict_left_mono [OF ec] by force - also have "\ \ e" using C - by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) - finally show ?thesis . - qed - show "\\<^sub>F n in at w. - \x\path_image \. - cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" - using twom [OF divide_pos_pos [OF \0 < e\ \C > 0\]] unfolding path_image_def - by (force intro: * elim: eventually_mono) - qed - show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" - by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto - have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) - \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" - by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto - have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = - (f u - f w) / (u - w) / k" - if "dist u w < d" for u - proof - - have u: "u \ s - path_image \" - by (metis subsetD d dist_commute mem_ball that) - 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) - done - qed - show ?thes2 - apply (simp add: has_field_derivative_iff del: power_Suc) - apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) - apply (simp add: \k \ 0\ **) - done -qed - -lemma Cauchy_next_derivative_circlepath: - assumes contf: "continuous_on (path_image (circlepath z r)) f" - and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" - and k: "k \ 0" - and w: "w \ ball z r" - shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" - (is "?thes2") -proof - - have "r > 0" using w - using ball_eq_empty by fastforce - have wim: "w \ ball z r - path_image (circlepath z r)" - using w by (auto simp: dist_norm) - show ?thes1 ?thes2 - by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; - auto simp: vector_derivative_circlepath norm_mult)+ -qed - - -text\ In particular, the first derivative formula.\ - -lemma Cauchy_derivative_integral_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" - (is "?thes2") -proof - - have [simp]: "r \ 0" using w - using ball_eq_empty by fastforce - have f: "continuous_on (path_image (circlepath z r)) f" - by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def) - have int: "\w. dist z w < r \ - ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" - by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) - show ?thes1 - apply (simp add: power2_eq_square) - apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) - apply (blast intro: int) - done - have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" - apply (simp add: power2_eq_square) - apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) - apply (blast intro: int) - done - then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" - by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) - show ?thes2 - by simp (rule fder) -qed - -subsection\Existence of all higher derivatives\ - -proposition derivative_is_holomorphic: - assumes "open S" - and fder: "\z. z \ S \ (f has_field_derivative f' z) (at z)" - shows "f' holomorphic_on S" -proof - - have *: "\h. (f' has_field_derivative h) (at z)" if "z \ S" for z - proof - - obtain r where "r > 0" and r: "cball z r \ S" - using open_contains_cball \z \ S\ \open S\ by blast - then have holf_cball: "f holomorphic_on cball z r" - apply (simp add: holomorphic_on_def) - using field_differentiable_at_within field_differentiable_def fder by blast - then have "continuous_on (path_image (circlepath z r)) f" - using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) - then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" - by (auto intro: continuous_intros)+ - have contf_cball: "continuous_on (cball z r) f" using holf_cball - by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) - have holf_ball: "f holomorphic_on ball z r" using holf_cball - using ball_subset_cball holomorphic_on_subset by blast - { fix w assume w: "w \ ball z r" - have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" - by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) - have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) - (at w)" - by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) - have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" - using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) - have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral - contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) - (circlepath z r)" - by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) - then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral - contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) - (circlepath z r)" - by (simp add: algebra_simps) - then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" - by (simp add: f'_eq) - } note * = this - show ?thesis - apply (rule exI) - apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) - apply (simp_all add: \0 < r\ * dist_norm) - done - qed - show ?thesis - by (simp add: holomorphic_on_open [OF \open S\] *) -qed - -lemma holomorphic_deriv [holomorphic_intros]: - "\f holomorphic_on S; open S\ \ (deriv f) holomorphic_on S" -by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) - -lemma analytic_deriv [analytic_intros]: "f analytic_on S \ (deriv f) analytic_on S" - using analytic_on_holomorphic holomorphic_deriv by auto - -lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv ^^ n) f holomorphic_on S" - by (induction n) (auto simp: holomorphic_deriv) - -lemma analytic_higher_deriv [analytic_intros]: "f analytic_on S \ (deriv ^^ n) f analytic_on S" - unfolding analytic_on_def using holomorphic_higher_deriv by blast - -lemma has_field_derivative_higher_deriv: - "\f holomorphic_on S; open S; x \ S\ - \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" -by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply - funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) - -lemma valid_path_compose_holomorphic: - assumes "valid_path g" and holo:"f holomorphic_on S" and "open S" "path_image g \ S" - shows "valid_path (f \ g)" -proof (rule valid_path_compose[OF \valid_path g\]) - fix x assume "x \ path_image g" - then show "f field_differentiable at x" - using analytic_on_imp_differentiable_at analytic_on_open assms holo by blast -next - have "deriv f holomorphic_on S" - using holomorphic_deriv holo \open S\ by auto - then show "continuous_on (path_image g) (deriv f)" - using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto -qed - - -subsection\Morera's theorem\ - -lemma Morera_local_triangle_ball: - assumes "\z. z \ S - \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ - (\b c. closed_segment b c \ ball a e - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0)" - shows "f analytic_on S" -proof - - { fix z assume "z \ S" - with assms obtain e a where - "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" - and 0: "\b c. closed_segment b c \ ball a e - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - by fastforce - 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 - qed (simp add: az) - } - then show ?thesis - by (simp add: analytic_on_def) -qed - -lemma Morera_local_triangle: - assumes "\z. z \ S - \ \t. open t \ z \ t \ continuous_on t f \ - (\a b c. convex hull {a,b,c} \ t - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0)" - shows "f analytic_on S" -proof - - { fix z assume "z \ S" - with assms obtain t where - "open t" and z: "z \ t" and contf: "continuous_on t f" - and 0: "\a b c. convex hull {a,b,c} \ t - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - by force - then obtain e where "e>0" and e: "ball z e \ t" - using open_contains_ball by blast - have [simp]: "continuous_on (ball z e) f" using contf - using continuous_on_subset e by blast - have eq0: "\b c. closed_segment b c \ ball z e \ - contour_integral (linepath z b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c z) f = 0" - by (meson 0 z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) - have "\e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ - (\b c. closed_segment b c \ ball a e \ - contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" - using \e > 0\ eq0 by force - } - then show ?thesis - by (simp add: Morera_local_triangle_ball) -qed - -proposition Morera_triangle: - "\continuous_on S f; open S; - \a b c. convex hull {a,b,c} \ S - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0\ - \ f analytic_on S" - using Morera_local_triangle by blast - -subsection\Combining theorems for higher derivatives including Leibniz rule\ - -lemma higher_deriv_linear [simp]: - "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" - by (induction n) auto - -lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" - by (induction n) auto - -lemma higher_deriv_ident [simp]: - "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" - apply (induction n, simp) - apply (metis higher_deriv_linear lambda_one) - done - -lemma higher_deriv_id [simp]: - "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" - by (simp add: id_def) - -lemma has_complex_derivative_funpow_1: - "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" - apply (induction n, auto) - apply (simp add: id_def) - by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) - -lemma higher_deriv_uminus: - assumes "f holomorphic_on S" "open S" and z: "z \ S" - shows "(deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - have "((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 - then show ?case - by (simp add: DERIV_imp_deriv) -qed - -lemma higher_deriv_add: - fixes z::complex - assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" - shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - have "((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 - 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" - 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 - -lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" - by (cases k) simp_all - -lemma higher_deriv_mult: - fixes z::complex - assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" - shows "(deriv ^^ n) (\w. f w * g w) z = - (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - "\n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - have sumeq: "(\i = 0..n. - of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) = - g z * deriv ((deriv ^^ n) f) z + (\i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" - apply (simp add: bb algebra_simps sum.distrib) - apply (subst (4) sum_Suc_reindex) - apply (auto simp: algebra_simps Suc_diff_le intro: sum.cong) - done - have "((deriv ^^ n) (\w. f w * g w) has_field_derivative - (\i = 0..Suc n. (Suc n choose i) * (deriv ^^ i) f z * (deriv ^^ (Suc n - i)) g z)) - (at z)" - apply (rule has_field_derivative_transform_within_open - [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) - 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 - then show ?case - unfolding funpow.simps o_apply - by (simp add: DERIV_imp_deriv) -qed - -lemma higher_deriv_transform_within_open: - fixes z::complex - assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" - and fg: "\w. w \ S \ f w = g w" - shows "(deriv ^^ i) f z = (deriv ^^ i) g z" -using z -by (induction i arbitrary: z) - (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) - -lemma higher_deriv_compose_linear: - fixes z::complex - assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \ S" - and fg: "\w. w \ S \ u * w \ T" - shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have holo0: "f holomorphic_on (*) u ` S" - by (meson fg f holomorphic_on_subset image_subset_iff) - have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S" - by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T) - have holo3: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S" - by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros) - have 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 "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 - 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 - also have "\ = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" - apply (subst complex_derivative_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 - finally show ?case - by simp -qed - -lemma higher_deriv_add_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_add show ?thesis - by (auto simp: analytic_at_two) -qed - -lemma higher_deriv_diff_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_diff show ?thesis - by (auto simp: analytic_at_two) -qed - -lemma higher_deriv_uminus_at: - "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" - using higher_deriv_uminus - by (auto simp: analytic_at) - -lemma higher_deriv_mult_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w * g w) z = - (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_mult show ?thesis - by (auto simp: analytic_at_two) -qed -text\ Nonexistence of isolated singularities and a stronger integral formula.\ - -proposition no_isolated_singularity: - fixes z::complex - assumes f: "continuous_on S f" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" - shows "f holomorphic_on S" -proof - - { fix z - assume "z \ S" and cdf: "\x. x \ S - K \ f field_differentiable at x" - have "f field_differentiable at z" - proof (cases "z \ K") - case False then show ?thesis by (blast intro: cdf \z \ S\) - next - case True - with finite_set_avoid [OF K, of z] - obtain d where "d>0" and d: "\x. \x\K; x \ z\ \ d \ dist z x" - by blast - obtain e where "e>0" and e: "ball z e \ S" - using S \z \ S\ by (force simp: open_contains_ball) - have fde: "continuous_on (ball z (min d e)) f" - by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) - have cont: "{a,b,c} \ ball z (min d e) \ continuous_on (convex hull {a, b, c}) f" for a b c - by (simp add: hull_minimal continuous_on_subset [OF fde]) - have fd: "\{a,b,c} \ ball z (min d e); x \ interior (convex hull {a, b, c}) - K\ - \ f field_differentiable at x" for a b c x - by (metis cdf Diff_iff Int_iff ball_min_Int subsetD convex_ball e interior_mono interior_subset subset_hull) - obtain g where "\w. w \ ball z (min d e) \ (g has_field_derivative f w) (at w within ball z (min d e))" - apply (rule contour_integral_convex_primitive - [OF convex_ball fde Cauchy_theorem_triangle_cofinite [OF _ K]]) - using cont fd by auto - then have "f holomorphic_on ball z (min d e)" - by (metis open_ball at_within_open derivative_is_holomorphic) - then show ?thesis - unfolding holomorphic_on_def - by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) - qed - } - with holf S K show ?thesis - by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) -qed - -lemma no_isolated_singularity': - fixes z::complex - assumes f: "\z. z \ K \ (f \ f z) (at z within S)" - and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" - shows "f holomorphic_on S" -proof (rule no_isolated_singularity[OF _ assms(2-)]) - show "continuous_on S f" unfolding continuous_on_def - proof - fix z assume z: "z \ S" - show "(f \ f z) (at z within S)" - proof (cases "z \ K") - case False - from holf have "continuous_on (S - K) f" - by (rule holomorphic_on_imp_continuous_on) - with z False have "(f \ f z) (at z within (S - K))" - by (simp add: continuous_on_def) - also from z K S False have "at z within (S - K) = at z within S" - by (subst (1 2) at_within_open) (auto intro: finite_imp_closed) - finally show "(f \ f z) (at z within S)" . - qed (insert assms z, simp_all) - qed -qed - -proposition Cauchy_integral_formula_convex: - assumes S: "convex S" and K: "finite K" and contf: "continuous_on S f" - and fcd: "(\x. x \ interior S - K \ f field_differentiable at x)" - and z: "z \ interior S" and vpg: "valid_path \" - and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - have *: "\x. x \ interior S \ f field_differentiable at x" - unfolding holomorphic_on_open [symmetric] field_differentiable_def - using no_isolated_singularity [where S = "interior S"] - by (meson K contf continuous_at_imp_continuous_on continuous_on_interior fcd - field_differentiable_at_within field_differentiable_def holomorphic_onI - holomorphic_on_imp_differentiable_at open_interior) - show ?thesis - by (rule Cauchy_integral_formula_weak [OF S finite.emptyI contf]) (use * assms in auto) -qed - -text\ Formula for higher derivatives.\ - -lemma Cauchy_has_contour_integral_higher_derivative_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) - (circlepath z r)" -using w -proof (induction k arbitrary: w) - case 0 then show ?case - using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) -next - case (Suc k) - have [simp]: "r > 0" using w - using ball_eq_empty by fastforce - have f: "continuous_on (path_image (circlepath z r)) f" - by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def less_imp_le) - obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" - using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] - by (auto simp: contour_integrable_on_def) - then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" - by (rule contour_integral_unique) - have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" - by (force simp: field_differentiable_def) - have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = - of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" - by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) - also have "\ = of_nat (Suc k) * X" - by (simp only: con) - finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . - then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" - by (metis deriv_cmult dnf_diff) - then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" - by (simp add: field_simps) - then show ?case - using of_nat_eq_0_iff X by fastforce -qed - -lemma Cauchy_higher_derivative_integral_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" - (is "?thes2") -proof - - have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) - (circlepath z r)" - using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] - by simp - show ?thes1 using * - using contour_integrable_on_def by blast - show ?thes2 - unfolding contour_integral_unique [OF *] by (simp add: field_split_simps) -qed - -corollary Cauchy_contour_integral_circlepath: - assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" - shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" -by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) - -lemma Cauchy_contour_integral_circlepath_2: - assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" - shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" - using Cauchy_contour_integral_circlepath [OF assms, of 1] - by (simp add: power2_eq_square) - - -subsection\A holomorphic function is analytic, i.e. has local power series\ - -theorem holomorphic_power_series: - assumes holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" -proof - - \ \Replacing \<^term>\r\ and the original (weak) premises with stronger ones\ - obtain r where "r > 0" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" - proof - have "cball z ((r + dist w z) / 2) \ ball z r" - using w by (simp add: dist_commute field_sum_of_halves subset_eq) - then show "f holomorphic_on cball z ((r + dist w z) / 2)" - by (rule holomorphic_on_subset [OF holf]) - have "r > 0" - using w by clarsimp (metis dist_norm le_less_trans norm_ge_zero) - then show "0 < (r + dist w z) / 2" - by simp (use zero_le_dist [of w z] in linarith) - qed (use w in \auto simp: dist_commute\) - then have holf: "f holomorphic_on ball z r" - using ball_subset_cball holomorphic_on_subset by blast - have contf: "continuous_on (cball z r) f" - by (simp add: holfc holomorphic_on_imp_continuous_on) - have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" - by (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) (simp add: \0 < r\) - obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" - by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) - obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" - and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" - proof - show "\u. cmod (u - z) = r \ r - dist z w \ cmod (u - w)" - by (metis add_diff_eq diff_add_cancel dist_norm norm_diff_ineq) - qed (use w in \auto simp: dist_norm norm_minus_commute\) - have ul: "uniform_limit (sphere z r) (\n x. (\kx. f x / (x - w)) sequentially" - unfolding uniform_limit_iff dist_norm - proof clarify - fix e::real - assume "0 < e" - have rr: "0 \ (r - k) / r" "(r - k) / r < 1" using k by auto - obtain n where n: "((r - k) / r) ^ n < e / B * k" - using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force - have "norm ((\k N" and r: "r = dist z u" for N u - proof - - have N: "((r - k) / r) ^ N < e / B * k" - apply (rule le_less_trans [OF power_decreasing n]) - using \n \ N\ k by auto - have u [simp]: "(u \ z) \ (u \ w)" - using \0 < r\ r w by auto - have wzu_not1: "(w - z) / (u - z) \ 1" - by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) - have "norm ((\kk = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)" - using \0 < B\ - apply (auto simp: geometric_sum [OF wzu_not1]) - apply (simp add: field_simps norm_mult [symmetric]) - done - also have "\ = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" - using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) - also have "\ = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" - by (simp add: algebra_simps) - also have "\ = norm (w - z) ^ N * norm (f u) / r ^ N" - by (simp add: norm_mult norm_power norm_minus_commute) - also have "\ \ (((r - k)/r)^N) * B" - using \0 < r\ w k - 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 - also have "\ < e * k" - using \0 < B\ N by (simp add: divide_simps) - also have "\ \ e * norm (u - w)" - using r kle \0 < e\ by (simp add: dist_commute dist_norm) - finally show ?thesis - by (simp add: field_split_simps norm_divide del: power_Suc) - qed - with \0 < r\ show "\\<^sub>F n in sequentially. \x\sphere z r. - norm ((\k\<^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) - done - have cic: "\u. (\y. \k0 < 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) - 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 - then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) - sums (2 * of_real pi * \ * f w)" - using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) - then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) - sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" - by (rule sums_divide) - then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) - sums f w" - by (simp add: field_simps) - then show ?thesis - by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) -qed - - -subsection\The Liouville theorem and the Fundamental Theorem of Algebra\ - -text\ These weak Liouville versions don't even need the derivative formula.\ - -lemma Liouville_weak_0: - assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" - shows "f z = 0" -proof (rule ccontr) - assume fz: "f z \ 0" - with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] - obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" - by (auto simp: dist_norm) - define R where "R = 1 + \B\ + norm z" - have "R > 0" unfolding R_def - proof - - have "0 \ cmod z + \B\" - by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) - then show "0 < 1 + \B\ + cmod z" - by linarith - qed - have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" - apply (rule Cauchy_integral_circlepath) - using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ - done - have "cmod (x - z) = R \ cmod (f x) * 2 < cmod (f z)" for x - unfolding R_def - by (rule B) (use norm_triangle_ineq4 [of x z] in auto) - with \R > 0\ fz show False - using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] - by (auto simp: less_imp_le norm_mult norm_divide field_split_simps) -qed - -proposition Liouville_weak: - assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" - shows "f z = l" - using Liouville_weak_0 [of "\z. f z - l"] - by (simp add: assms holomorphic_on_diff LIM_zero) - -proposition Liouville_weak_inverse: - assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" - obtains z where "f z = 0" -proof - - { assume f: "\z. f z \ 0" - have 1: "(\x. 1 / f x) holomorphic_on UNIV" - by (simp add: holomorphic_on_divide assms f) - have 2: "((\x. 1 / f x) \ 0) at_infinity" - 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 - have False - using Liouville_weak_0 [OF 1 2] f by simp - } - then show ?thesis - using that by blast -qed - -text\ In particular we get the Fundamental Theorem of Algebra.\ - -theorem fundamental_theorem_of_algebra: - fixes a :: "nat \ complex" - assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" - obtains z where "(\i\n. a i * z^i) = 0" -using assms -proof (elim disjE bexE) - assume "a 0 = 0" then show ?thesis - by (auto simp: that [of 0]) -next - fix i - assume i: "i \ {1..n}" and nz: "a i \ 0" - have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" - by (rule holomorphic_intros)+ - show thesis - proof (rule Liouville_weak_inverse [OF 1]) - show "\\<^sub>F x in at_infinity. B \ cmod (\i\n. a i * x ^ i)" for B - using i polyfun_extremal nz by force - qed (use that in auto) -qed - -subsection\Weierstrass convergence theorem\ - -lemma holomorphic_uniform_limit: - assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" - and ulim: "uniform_limit (cball z r) f g F" - and F: "\ trivial_limit F" - obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" -proof (cases r "0::real" rule: linorder_cases) - case less then show ?thesis by (force simp: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) -next - case equal then show ?thesis - by (force simp: holomorphic_on_def intro: that) -next - case greater - have contg: "continuous_on (cball z r) g" - using cont uniform_limit_theorem [OF eventually_mono ulim F] by blast - have "path_image (circlepath z r) \ cball z r" - using \0 < r\ by auto - then have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" - by (intro continuous_intros continuous_on_subset [OF contg]) - have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" - if w: "w \ ball z r" for w - proof - - define d where "d = (r - norm(w - z))" - have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) - have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" - unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) - have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" - 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) - 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 - have g_cint: "(\u. g u/(u - w)) contour_integrable_on circlepath z r" - by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) - have cif_tends_cig: "((\n. contour_integral(circlepath z r) (\u. f n u / (u - w))) \ contour_integral(circlepath z r) (\u. g u/(u - w))) F" - by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) - have f_tends_cig: "((\n. 2 * of_real pi * \ * f n w) \ contour_integral (circlepath z r) (\u. g u / (u - w))) F" - proof (rule Lim_transform_eventually) - show "\\<^sub>F x in F. contour_integral (circlepath z r) (\u. f x u / (u - w)) - = 2 * of_real pi * \ * f x w" - 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 fastforce - then have "((\u. g u / (2 * of_real pi * \ * (u - w))) has_contour_integral g w) (circlepath z r)" - using has_contour_integral_div [where c = "2 * of_real pi * \"] - by (force simp: field_simps) - then show ?thesis - by (simp add: dist_norm) - qed - show ?thesis - using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] - by (fastforce simp add: holomorphic_on_open contg intro: that) -qed - - -text\ Version showing that the limit is the limit of the derivatives.\ - -proposition has_complex_derivative_uniform_limit: - fixes z::complex - assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ - (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" - and ulim: "uniform_limit (cball z r) f g F" - and F: "\ trivial_limit F" and "0 < r" - obtains g' where - "continuous_on (cball z r) g" - "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" -proof - - let ?conint = "contour_integral (circlepath z r)" - have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" - by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] ulim F]; - auto simp: holomorphic_on_open field_differentiable_def)+ - then obtain g' where g': "\x. x \ ball z r \ (g has_field_derivative g' x) (at x)" - using DERIV_deriv_iff_has_field_derivative - by (fastforce simp add: holomorphic_on_open) - then have derg: "\x. x \ ball z r \ deriv g x = g' x" - by (simp add: DERIV_imp_deriv) - have tends_f'n_g': "((\n. f' n w) \ g' w) F" if w: "w \ ball z r" for w - proof - - have eq_f': "?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \)" - if cont_fn: "continuous_on (cball z r) (f n)" - and fnd: "\w. w \ ball z r \ (f n has_field_derivative f' n w) (at w)" for n - proof - - have hol_fn: "f n holomorphic_on ball z r" - using fnd by (force simp: holomorphic_on_open) - have "(f n has_field_derivative 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)) (at w)" - by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) - then have f': "f' n w = 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)" - using DERIV_unique [OF fnd] w by blast - show ?thesis - by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] field_split_simps) - qed - define d where "d = (r - norm(w - z))^2" - have "d > 0" - using w by (simp add: dist_commute dist_norm d_def) - have dle: "d \ cmod ((y - w)\<^sup>2)" if "r = cmod (z - y)" for y - proof - - have "w \ ball z (cmod (z - y))" - using that w by fastforce - then have "cmod (w - z) \ cmod (z - y)" - by (simp add: dist_complex_def norm_minus_commute) - moreover have "cmod (z - y) - cmod (w - z) \ cmod (y - w)" - by (metis diff_add_cancel diff_add_eq_diff_diff_swap norm_minus_commute norm_triangle_ineq2) - ultimately show ?thesis - using that by (simp add: d_def norm_power power_mono) - qed - have 1: "\\<^sub>F n in F. (\x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" - by (force simp: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) - have 2: "uniform_limit (sphere z r) (\n x. f n x / (x - w)\<^sup>2) (\x. g x / (x - w)\<^sup>2) F" - unfolding uniform_limit_iff - proof clarify - fix e::real - assume "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 - qed - have "((\n. contour_integral (circlepath z r) (\x. f n x / (x - w)\<^sup>2)) - \ contour_integral (circlepath z r) ((\x. g x / (x - w)\<^sup>2))) F" - by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \0 < r\]) - then have tendsto_0: "((\n. 1 / (2 * of_real pi * \) * (?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2))) \ 0) F" - using Lim_null by (force intro!: tendsto_mult_right_zero) - have "((\n. f' n w - g' w) \ 0) F" - apply (rule Lim_transform_eventually [OF tendsto_0]) - apply (force simp: divide_simps intro: eq_f' eventually_mono [OF cont]) - done - then show ?thesis using Lim_null by blast - qed - obtain g' where "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" - by (blast intro: tends_f'n_g' g') - then show ?thesis using g - using that by blast -qed - - -subsection\<^marker>\tag unimportant\ \Some more simple/convenient versions for applications\ - -lemma holomorphic_uniform_sequence: - assumes S: "open S" - and hol_fn: "\n. (f n) holomorphic_on S" - and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" - shows "g holomorphic_on S" -proof - - have "\f'. (g has_field_derivative f') (at z)" if "z \ S" for z - proof - - obtain r where "0 < r" and r: "cball z r \ S" - and ul: "uniform_limit (cball z r) f g sequentially" - using ulim_g [OF \z \ S\] by blast - have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ f n holomorphic_on ball z r" - proof (intro eventuallyI conjI) - show "continuous_on (cball z r) (f x)" for x - using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r by blast - show "f x holomorphic_on ball z r" for x - by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) - qed - show ?thesis - apply (rule holomorphic_uniform_limit [OF *]) - using \0 < r\ centre_in_ball ul - apply (auto simp: holomorphic_on_open) - done - qed - with S show ?thesis - by (simp add: holomorphic_on_open) -qed - -lemma has_complex_derivative_uniform_sequence: - fixes S :: "complex set" - assumes S: "open S" - and hfd: "\n x. x \ S \ ((f n) has_field_derivative f' n x) (at x)" - and ulim_g: "\x. x \ S - \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" - shows "\g'. \x \ S. (g has_field_derivative g' x) (at x) \ ((\n. f' n x) \ g' x) sequentially" -proof - - have y: "\y. (g has_field_derivative y) (at z) \ (\n. f' n z) \ y" if "z \ S" for z - proof - - obtain r where "0 < r" and r: "cball z r \ S" - and ul: "uniform_limit (cball z r) f g sequentially" - using ulim_g [OF \z \ S\] by blast - have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ - (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))" - proof (intro eventuallyI conjI ballI) - show "continuous_on (cball z r) (f x)" for x - by (meson S continuous_on_subset hfd holomorphic_on_imp_continuous_on holomorphic_on_open r) - show "w \ ball z r \ (f x has_field_derivative f' x w) (at w)" for w x - using ball_subset_cball hfd r by blast - qed - show ?thesis - by (rule has_complex_derivative_uniform_limit [OF *, of g]) (use \0 < r\ ul in \force+\) - qed - show ?thesis - by (rule bchoice) (blast intro: y) -qed - -subsection\On analytic functions defined by a series\ - -lemma series_and_derivative_comparison: - fixes S :: "complex set" - assumes S: "open S" - and h: "summable h" - and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\\<^sub>F n in sequentially. \x\S. norm (f n x) \ h n" - obtains g g' where "\x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -proof - - obtain g where g: "uniform_limit S (\n x. \id>0. cball x d \ S \ uniform_limit (cball x d) (\n x. \i S" for x - proof - - obtain d where "d>0" and d: "cball x d \ S" - using open_contains_cball [of "S"] \x \ S\ S by blast - show ?thesis - proof (intro conjI exI) - show "uniform_limit (cball x d) (\n x. \id > 0\ d in auto) - qed - have "\x. x \ S \ (\n. \i g x" - by (metis tendsto_uniform_limitI [OF g]) - moreover have "\g'. \x\S. (g has_field_derivative g' x) (at x) \ (\n. \i g' x" - by (rule has_complex_derivative_uniform_sequence [OF S]) (auto intro: * hfd DERIV_sum)+ - ultimately show ?thesis - by (metis sums_def that) -qed - -text\A version where we only have local uniform/comparative convergence.\ - -lemma series_and_derivative_comparison_local: - fixes S :: "complex set" - assumes S: "open S" - and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ (\\<^sub>F n in sequentially. \y\ball x d \ S. norm (f n y) \ h n)" - shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -proof - - have "\y. (\n. f n z) sums (\n. f n z) \ (\n. f' n z) sums y \ ((\x. \n. f n x) has_field_derivative y) (at z)" - if "z \ S" for z - proof - - obtain d h where "0 < d" "summable h" and le_h: "\\<^sub>F n in sequentially. \y\ball z d \ S. norm (f n y) \ h n" - using to_g \z \ S\ by meson - then obtain r where "r>0" and r: "ball z r \ ball z d \ S" using \z \ S\ S - by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) - have 1: "open (ball z d \ S)" - by (simp add: open_Int S) - have 2: "\n x. x \ ball z d \ S \ (f n has_field_derivative f' n x) (at x)" - by (auto simp: hfd) - obtain g g' where gg': "\x \ ball z d \ S. ((\n. f n x) sums g x) \ - ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" - by (auto intro: le_h series_and_derivative_comparison [OF 1 \summable h\ hfd]) - then have "(\n. f' n z) sums g' z" - by (meson \0 < r\ centre_in_ball contra_subsetD r) - moreover have "(\n. f n z) sums (\n. f n z)" - using summable_sums centre_in_ball \0 < d\ \summable h\ le_h - by (metis (full_types) Int_iff gg' summable_def that) - moreover have "((\x. \n. f n x) has_field_derivative g' z) (at z)" - proof (rule has_field_derivative_transform_within) - show "\x. dist x z < r \ g x = (\n. f n x)" - by (metis subsetD dist_commute gg' mem_ball r sums_unique) - qed (use \0 < r\ gg' \z \ S\ \0 < d\ in auto) - ultimately show ?thesis by auto - qed - then show ?thesis - by (rule_tac x="\x. suminf (\n. f n x)" in exI) meson -qed - - -text\Sometimes convenient to compare with a complex series of positive reals. (?)\ - -lemma series_and_derivative_comparison_complex: - fixes S :: "complex set" - assumes S: "open S" - and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" - shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -apply (rule series_and_derivative_comparison_local [OF S hfd], assumption) -apply (rule ex_forward [OF to_g], assumption) -apply (erule exE) -apply (rule_tac x="Re \ h" in exI) -apply (force simp: summable_Re o_def nonneg_Reals_cmod_eq_Re image_subset_iff) -done - -text\Sometimes convenient to compare with a complex series of positive reals. (?)\ -lemma series_differentiable_comparison_complex: - fixes S :: "complex set" - assumes S: "open S" - and hfd: "\n x. x \ S \ f n field_differentiable (at x)" - and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" - obtains g where "\x \ S. ((\n. f n x) sums g x) \ g field_differentiable (at x)" -proof - - have hfd': "\n x. x \ S \ (f n has_field_derivative deriv (f n) x) (at x)" - using hfd field_differentiable_derivI by blast - have "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. deriv (f n) x) sums g' x) \ (g has_field_derivative g' x) (at x)" - by (metis series_and_derivative_comparison_complex [OF S hfd' to_g]) - then show ?thesis - using field_differentiable_def that by blast -qed - -text\In particular, a power series is analytic inside circle of convergence.\ - -lemma power_series_and_derivative_0: - fixes a :: "nat \ complex" and r::real - assumes "summable (\n. a n * r^n)" - shows "\g g'. \z. cmod z < r \ - ((\n. a n * z^n) sums g z) \ ((\n. of_nat n * a n * z^(n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" -proof (cases "0 < r") - case True - have der: "\n z. ((\x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)" - by (rule derivative_eq_intros | simp)+ - have y_le: "\cmod (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 "summable (\n. a n * complex_of_real r ^ n)" - using assms \r > 0\ by simp - moreover have "\z. cmod z < r \ cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)" - using \r > 0\ - by (simp flip: of_real_add) - ultimately have sum: "\z. cmod z < r \ summable (\n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)" - by (rule power_series_conv_imp_absconv_weak) - have "\g g'. \z \ ball 0 r. (\n. (a n) * z ^ n) sums g z \ - (\n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \ (g has_field_derivative g' z) (at z)" - apply (rule series_and_derivative_comparison_complex [OF open_ball der]) - apply (rule_tac x="(r - norm z)/2" in exI) - apply (rule_tac x="\n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI) - using \r > 0\ - apply (auto simp: sum eventually_sequentially norm_mult norm_power dist_norm intro!: mult_left_mono power_mono y_le) - done - then show ?thesis - by (simp add: ball_def) -next - case False then show ?thesis - apply (simp add: not_less) - using less_le_trans norm_not_less_zero by blast -qed - -proposition\<^marker>\tag unimportant\ power_series_and_derivative: - fixes a :: "nat \ complex" and r::real - assumes "summable (\n. a n * r^n)" - obtains g g' where "\z \ ball w r. - ((\n. a n * (z - w) ^ n) sums g z) \ ((\n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \ - (g has_field_derivative g' z) (at z)" - using power_series_and_derivative_0 [OF assms] - apply clarify - apply (rule_tac g="(\z. g(z - w))" in that) - using DERIV_shift [where z="-w"] - apply (auto simp: norm_minus_commute Ball_def dist_norm) - done - -proposition\<^marker>\tag unimportant\ power_series_holomorphic: - assumes "\w. w \ ball z r \ ((\n. a n*(w - z)^n) sums f w)" - shows "f holomorphic_on ball z r" -proof - - have "\f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w - proof - - have inb: "z + complex_of_real ((dist z w + r) / 2) \ ball z r" - proof - - have wz: "cmod (w - z) < r" using w - by (auto simp: field_split_simps dist_norm norm_minus_commute) - then have "0 \ r" - by (meson less_eq_real_def norm_ge_zero order_trans) - show ?thesis - using w by (simp add: dist_norm \0\r\ flip: of_real_add) - qed - have sum: "summable (\n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))" - using assms [OF inb] by (force simp: summable_def dist_norm) - obtain g g' where gg': "\u. u \ ball z ((cmod (z - w) + r) / 2) \ - (\n. a n * (u - z) ^ n) sums g u \ - (\n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \ (g has_field_derivative g' u) (at u)" - by (rule power_series_and_derivative [OF sum, of z]) fastforce - have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u - proof - - have less: "cmod (z - u) * 2 < cmod (z - w) + r" - using that dist_triangle2 [of z u w] - by (simp add: dist_norm [symmetric] algebra_simps) - 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 - qed - have "(f has_field_derivative g' w) (at w)" - by (rule has_field_derivative_transform_within [where d="(r - norm(z - w))/2"]) - (use w gg' [of w] in \(force simp: dist_norm)+\) - then show ?thesis .. - qed - then show ?thesis by (simp add: holomorphic_on_open) -qed - -corollary holomorphic_iff_power_series: - "f holomorphic_on ball z r \ - (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" - apply (intro iffI ballI holomorphic_power_series, assumption+) - apply (force intro: power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) - done - -lemma power_series_analytic: - "(\w. w \ ball z r \ (\n. a n*(w - z)^n) sums f w) \ f analytic_on ball z r" - by (force simp: analytic_on_open intro!: power_series_holomorphic) - -lemma analytic_iff_power_series: - "f analytic_on ball z r \ - (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" - by (simp add: analytic_on_open holomorphic_iff_power_series) - -subsection\<^marker>\tag unimportant\ \Equality between holomorphic functions, on open ball then connected set\ - -lemma holomorphic_fun_eq_on_ball: - "\f holomorphic_on ball z r; g holomorphic_on ball z r; - w \ ball z r; - \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ - \ f w = g w" - apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) - apply (auto simp: holomorphic_iff_power_series) - done - -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 - -lemma holomorphic_fun_eq_0_on_connected: - assumes holf: "f holomorphic_on S" and "open S" - and cons: "connected S" - and der: "\n. (deriv ^^ n) f z = 0" - and "z \ S" "w \ S" - shows "f w = 0" -proof - - have *: "ball x e \ (\n. {w \ S. (deriv ^^ n) f w = 0})" - if "\u. (deriv ^^ u) f x = 0" "ball x e \ S" for x e - proof - - have "\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 (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" - apply (rule open_subset, force) - 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})" - 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" - using \e>0\ e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) - 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 -qed - -lemma holomorphic_fun_eq_on_connected: - assumes "f holomorphic_on S" "g holomorphic_on S" and "open S" "connected S" - and "\n. (deriv ^^ n) f z = (deriv ^^ n) g z" - and "z \ S" "w \ S" - shows "f w = g w" -proof (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" S z, simplified]) - show "(\x. f x - g x) holomorphic_on S" - by (intro assms holomorphic_intros) - show "\n. (deriv ^^ n) (\x. f x - g x) z = 0" - using assms higher_deriv_diff by auto -qed (use assms in auto) - -lemma holomorphic_fun_eq_const_on_connected: - assumes holf: "f holomorphic_on S" and "open S" - and cons: "connected S" - and der: "\n. 0 < n \ (deriv ^^ n) f z = 0" - and "z \ S" "w \ S" - shows "f w = f z" -proof (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" S z, simplified]) - show "(\w. f w - f z) holomorphic_on S" - by (intro assms holomorphic_intros) - show "\n. (deriv ^^ n) (\w. f w - f z) z = 0" - by (subst higher_deriv_diff) (use assms in \auto intro: holomorphic_intros\) -qed (use assms in auto) - -subsection\<^marker>\tag unimportant\ \Some basic lemmas about poles/singularities\ - -lemma pole_lemma: - assumes holf: "f holomorphic_on S" and a: "a \ interior S" - shows "(\z. if z = a then deriv f a - else (f z - f a) / (z - a)) holomorphic_on S" (is "?F holomorphic_on S") -proof - - have F1: "?F field_differentiable (at u within S)" if "u \ S" "u \ a" for u - proof - - have fcd: "f field_differentiable at u within S" - using holf holomorphic_on_def by (simp add: \u \ S\) - have cd: "(\z. (f z - f a) / (z - a)) field_differentiable at u within S" - by (rule fcd derivative_intros | simp add: that)+ - have "0 < dist a u" using that dist_nz by blast - then show ?thesis - by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \u \ S\) - qed - have F2: "?F field_differentiable at a" if "0 < e" "ball a e \ S" for e - proof - - have holfb: "f holomorphic_on ball a e" - by (rule holomorphic_on_subset [OF holf \ball a e \ S\]) - have 2: "?F holomorphic_on ball a e - {a}" - 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 have "f field_differentiable at a" - using holfb \0 < e\ holomorphic_on_imp_differentiable_at by auto - with True show ?thesis - by (auto simp: continuous_at has_field_derivative_iff simp flip: DERIV_deriv_iff_field_differentiable - elim: rev_iffD1 [OF _ LIM_equal]) - next - case False with 2 that show ?thesis - by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) - qed - then have 1: "continuous_on (ball a e) ?F" - by (clarsimp simp: continuous_on_eq_continuous_at) - have "?F holomorphic_on ball a e" - by (auto intro: no_isolated_singularity [OF 1 2]) - with that show ?thesis - by (simp add: holomorphic_on_open field_differentiable_def [symmetric] - field_differentiable_at_within) - qed - show ?thesis - proof - fix x assume "x \ S" show "?F field_differentiable at x within S" - proof (cases "x=a") - case True then show ?thesis - using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) - next - case False with F1 \x \ S\ - show ?thesis by blast - qed - qed -qed - -lemma pole_theorem: - assumes holg: "g holomorphic_on S" and a: "a \ interior S" - and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) holomorphic_on S" - using pole_lemma [OF holg a] - by (rule holomorphic_transform) (simp add: eq field_split_simps) - -lemma pole_lemma_open: - assumes "f holomorphic_on S" "open S" - shows "(\z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on S" -proof (cases "a \ S") - case True with assms interior_eq pole_lemma - show ?thesis by fastforce -next - case False with assms show ?thesis - apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) - apply (rule field_differentiable_transform_within [where f = "\z. (f z - f a)/(z - a)" and d = 1]) - apply (rule derivative_intros | force)+ - done -qed - -lemma pole_theorem_open: - assumes holg: "g holomorphic_on S" and S: "open S" - and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) holomorphic_on S" - using pole_lemma_open [OF holg S] - by (rule holomorphic_transform) (auto simp: eq divide_simps) - -lemma pole_theorem_0: - assumes holg: "g holomorphic_on S" and a: "a \ interior S" - and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f holomorphic_on S" - using pole_theorem [OF holg a eq] - by (rule holomorphic_transform) (auto simp: eq field_split_simps) - -lemma pole_theorem_open_0: - assumes holg: "g holomorphic_on S" and S: "open S" - and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f holomorphic_on S" - using pole_theorem_open [OF holg S eq] - by (rule holomorphic_transform) (auto simp: eq field_split_simps) - -lemma pole_theorem_analytic: - assumes g: "g analytic_on S" - and eq: "\z. z \ S - \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" - shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" (is "?F analytic_on S") - unfolding analytic_on_def -proof - fix x - assume "x \ S" - with g obtain e where "0 < e" and e: "g holomorphic_on ball x e" - by (auto simp add: analytic_on_def) - obtain d where "0 < d" and d: "\w. w \ ball x d - {a} \ g w = (w - a) * f w" - using \x \ S\ eq by blast - have "?F holomorphic_on ball x (min d e)" - using d e \x \ S\ by (fastforce simp: holomorphic_on_subset subset_ball intro!: pole_theorem_open) - then show "\e>0. ?F holomorphic_on ball x e" - using \0 < d\ \0 < e\ not_le by fastforce -qed - -lemma pole_theorem_analytic_0: - assumes g: "g analytic_on S" - and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f analytic_on S" -proof - - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" - by auto - show ?thesis - using pole_theorem_analytic [OF g eq] by simp -qed - -lemma pole_theorem_analytic_open_superset: - assumes g: "g analytic_on S" and "S \ T" "open T" - and eq: "\z. z \ T - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) analytic_on S" -proof (rule pole_theorem_analytic [OF g]) - fix z - assume "z \ S" - then obtain e where "0 < e" and e: "ball z e \ T" - using assms openE by blast - then show "\d>0. \w\ball z d - {a}. g w = (w - a) * f w" - using eq by auto -qed - -lemma pole_theorem_analytic_open_superset_0: - assumes g: "g analytic_on S" "S \ T" "open T" "\z. z \ T - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f analytic_on S" -proof - - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" - by auto - have "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" - by (rule pole_theorem_analytic_open_superset [OF g]) - then show ?thesis by simp -qed - - -subsection\General, homology form of Cauchy's theorem\ - -text\Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\ - -lemma contour_integral_continuous_on_linepath_2D: - assumes "open U" and cont_dw: "\w. w \ U \ F w contour_integrable_on (linepath a b)" - and cond_uu: "continuous_on (U \ U) (\(x,y). F x y)" - and abu: "closed_segment a b \ U" - shows "continuous_on U (\w. contour_integral (linepath a b) (F w))" -proof - - have *: "\d>0. \x'\U. dist x' w < d \ - dist (contour_integral (linepath a b) (F x')) - (contour_integral (linepath a b) (F w)) \ \" - if "w \ U" "0 < \" "a \ b" for w \ - proof - - obtain \ where "\>0" and \: "cball w \ \ U" using open_contains_cball \open U\ \w \ U\ by force - let ?TZ = "cball w \ \ closed_segment a b" - have "uniformly_continuous_on ?TZ (\(x,y). F x y)" - proof (rule compact_uniformly_continuous) - show "continuous_on ?TZ (\(x,y). F x y)" - by (rule continuous_on_subset[OF cond_uu]) (use SigmaE \ abu in blast) - show "compact ?TZ" - by (simp add: compact_Times) - qed - then obtain \ where "\>0" - and \: "\x x'. \x\?TZ; x'\?TZ; dist x' x < \\ \ - dist ((\(x,y). F x y) x') ((\(x,y). F x y) x) < \/norm(b - a)" - apply (rule uniformly_continuous_onE [where e = "\/norm(b - a)"]) - using \0 < \\ \a \ b\ by auto - have \: "\norm (w - x1) \ \; x2 \ closed_segment a b; - norm (w - x1') \ \; x2' \ closed_segment a b; norm ((x1', x2') - (x1, x2)) < \\ - \ norm (F x1' x2' - F x1 x2) \ \ / cmod (b - a)" - for x1 x2 x1' x2' - using \ [of "(x1,x2)" "(x1',x2')"] by (force simp: dist_norm) - have le_ee: "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \" - if "x' \ U" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' - proof - - have "(\x. F x' x - F w x) contour_integrable_on linepath a b" - by (simp add: \w \ U\ cont_dw contour_integrable_diff that) - then have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" - apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \]) - using \0 < \\ \0 < \\ that apply (auto simp: norm_minus_commute) - done - 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 - 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 - -text\This version has \<^term>\polynomial_function \\ as an additional assumption.\ -lemma Cauchy_integral_formula_global_weak: - assumes "open U" and holf: "f holomorphic_on U" - and z: "z \ U" and \: "polynomial_function \" - and pasz: "path_image \ \ U - {z}" and loop: "pathfinish \ = pathstart \" - and zero: "\w. w \ U \ winding_number \ w = 0" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - obtain \' where pf\': "polynomial_function \'" and \': "\x. (\ has_vector_derivative (\' x)) (at x)" - using has_vector_derivative_polynomial_function [OF \] by blast - then have "bounded(path_image \')" - by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) - then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" - using bounded_pos by force - define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w - define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" - have "path \" "valid_path \" using \ - by (auto simp: path_polynomial_function valid_path_polynomial_function) - then have ov: "open v" - by (simp add: v_def open_winding_number_levelsets loop) - have uv_Un: "U \ v = UNIV" - using pasz zero by (auto simp: v_def) - have conf: "continuous_on U f" - by (metis holf holomorphic_on_imp_continuous_on) - have hol_d: "(d y) holomorphic_on U" if "y \ U" for y - proof - - have *: "(\c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on U" - by (simp add: holf pole_lemma_open \open U\) - then have "isCont (\x. if x = y then deriv f y else (f x - f y) / (x - y)) y" - using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that \open U\ by fastforce - then have "continuous_on U (d y)" - 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\) - 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: \open U\) - qed - have cint_fxy: "(\x. (f x - f y) / (x - y)) contour_integrable_on \" if "y \ path_image \" for y - proof (rule contour_integrable_holomorphic_simple [where S = "U-{y}"]) - show "(\x. (f x - f y) / (x - y)) holomorphic_on U - {y}" - by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) - show "path_image \ \ U - {y}" - using pasz that by blast - qed (auto simp: \open U\ open_delete \valid_path \\) - define h where - "h z = (if z \ U then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z - have U: "((d z) has_contour_integral h z) \" if "z \ U" for z - proof - - have "d z holomorphic_on U" - by (simp add: hol_d that) - with that show ?thesis - apply (simp add: h_def) - 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" - using v_def z by auto - then have "((\x. 1 / (x - z)) has_contour_integral 0) \" - using z v_def has_contour_integral_winding_number [OF \valid_path \\] by fastforce - then have "((\x. f z * (1 / (x - z))) has_contour_integral 0) \" - using has_contour_integral_lmul by fastforce - then have "((\x. f z / (x - z)) has_contour_integral 0) \" - by (simp add: field_split_simps) - moreover have "((\x. (f x - f z) / (x - z)) has_contour_integral contour_integral \ (d z)) \" - using z - apply (auto simp: v_def) - apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) - done - ultimately have *: "((\x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \ (d z))) \" - by (rule has_contour_integral_add) - have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (d z)) \" - if "z \ U" - using * by (auto simp: divide_simps has_contour_integral_eq) - moreover have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (\w. f w / (w - z))) \" - if "z \ U" - 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 - ultimately show ?thesis - using z by (simp add: h_def) - qed - have znot: "z \ path_image \" - using pasz by blast - obtain d0 where "d0>0" and d0: "\x y. x \ path_image \ \ y \ - U \ d0 \ dist x y" - using separate_compact_closed [of "path_image \" "-U"] pasz \open U\ - 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" - apply (rule that [of "d0/2"]) - using \0 < d0\ - apply (auto simp: dist_norm dest: d0) - done - 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\ - 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 - obtain L where "L>0" - and L: "\f B. \f holomorphic_on interior T; \z. z\interior T \ cmod (f z) \ B\ \ - cmod (contour_integral \ f) \ L * B" - using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] - by blast - have "bounded(f ` T)" - by (meson \compact T\ compact_continuous_image compact_imp_bounded conf continuous_on_subset T) - then obtain D where "D>0" and D: "\x. x \ T \ norm (f x) \ D" - by (auto simp: bounded_pos) - obtain C where "C>0" and C: "\x. x \ T \ norm x \ C" - using \compact T\ bounded_pos compact_imp_bounded by force - have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y - proof - - have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp - with le have ybig: "norm y > C" by force - with C have "y \ T" by force - then have ynot: "y \ path_image \" - using subt interior_subset by blast - have [simp]: "winding_number \ y = 0" - 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 - 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 - have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior T" for z - proof - - have "D * L / e + cmod z \ cmod y" - using le C [of z] z using interior_subset by force - then have DL2: "D * L / e \ cmod (z - y)" - using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) - have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" - by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) - also have "\ \ D * (e / L / D)" - 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 - finally show ?thesis . - qed - have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" - by (simp add: dist_norm) - also have "\ \ L * (D * (e / L / D))" - by (rule L [OF holint leD]) - also have "\ = e" - using \L>0\ \0 < D\ by auto - finally show ?thesis . - qed - then have "(h \ 0) at_infinity" - by (meson Lim_at_infinityI) - moreover have "h holomorphic_on UNIV" - proof - - have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" - if "x \ U" "z \ U" "x \ z" for x z - using that conf - apply (simp add: split_def continuous_on_eq_continuous_at \open U\) - apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ - done - have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" - by (rule continuous_intros)+ - have open_uu_Id: "open (U \ U - Id)" - 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 - 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 - proof (rule Lim_withinI) - fix e::real assume "0 < e" - obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" - using \0 < e\ continuous_within_E [OF con_derf [OF \x \ U\]] - by (metis UNIV_I dist_norm) - obtain k2 where "k2>0" and k2: "ball x k2 \ U" - by (blast intro: openE [OF \open U\] \x \ U\) - have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" - if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" - for x' z' - proof - - have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w - apply (drule segment_furthest_le [where y=x]) - by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) - have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w - by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) - have f_has_der: "\x. x \ U \ (f has_field_derivative deriv f x) (at x within U)" - by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \open U\) - have "closed_segment x' z' \ U" - by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) - then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" - using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp - then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" - by (rule has_contour_integral_div) - have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" - apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) - using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] - \e > 0\ \z' \ x'\ - apply (auto simp: norm_divide divide_simps derf_le) - done - also have "\ \ e" using \0 < e\ by simp - finally show ?thesis . - qed - show "\d>0. \xa\U \ U. - 0 < dist xa (x, x) \ dist xa (x, x) < d \ - dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" - apply (rule_tac x="min k1 k2" in exI) - using \k1>0\ \k2>0\ \e>0\ - apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) - 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" - 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)" - apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') - apply (simp add: tendsto_within_open_NO_MATCH open_Times \open U\, clarify) - apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) - using con_ff - apply (auto simp: continuous_within) - done - have hol_dw: "(\z. d z w) holomorphic_on U" if "w \ U" for w - proof - - have "continuous_on U ((\(x,y). d x y) \ (\z. (w,z)))" - by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ - then have *: "continuous_on U (\z. if w = z then deriv f z else (f w - f z) / (w - z))" - by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) - have **: "\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 - 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 - 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 {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) \ \)" - proof (rule continuous_on_compose) - show "continuous_on {0..1} \" - using \path \\ path_def by blast - show "continuous_on (\ ` {0..1}) (\w. contour_integral (linepath a b) (\z. d z w))" - using pasz unfolding path_image_def - by (auto intro!: continuous_on_subset [OF cont_cint_d]) - qed - have 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 \']) - 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) - finally have cint_h_eq: - "contour_integral (linepath a b) h = - contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . - note cint_cint cint_h_eq - } note cint_h = this - have conthu: "continuous_on U h" - proof (simp add: continuous_on_sequentially, clarify) - fix a x - assume x: "x \ U" and au: "\n. a n \ U" and ax: "a \ x" - then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" - by (meson U contour_integrable_on_def eventuallyI) - obtain dd where "dd>0" and dd: "cball x dd \ U" using open_contains_cball \open U\ x by force - have A2: "uniform_limit (path_image \) (\n. d (a n)) (d x) sequentially" - unfolding uniform_limit_iff dist_norm - proof clarify - fix ee::real - assume "0 < ee" - show "\\<^sub>F n in sequentially. \\\path_image \. cmod (d (a n) \ - d x \) < ee" - proof - - let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" - have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" - 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 - then obtain kk where "kk>0" - and kk: "\x x'. \x \ ?ddpa; x' \ ?ddpa; dist x' x < kk\ \ - dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" - by (rule uniformly_continuous_onE [where e = ee]) (use \0 < ee\ in auto) - have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" - for w z - using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm) - 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 - qed - qed - have "(\n. contour_integral \ (d (a n))) \ contour_integral \ (d x)" - by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \valid_path \\) - then have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" - by (simp add: h_def x) - then show "(h \ a) \ h x" - by (simp add: h_def x au o_def) - qed - show ?thesis - proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) - fix z0 - consider "z0 \ v" | "z0 \ U" using uv_Un by blast - then show "h field_differentiable at z0" - proof cases - assume "z0 \ v" then show ?thesis - using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ - by (auto simp: field_differentiable_def v_def) - next - assume "z0 \ U" then - obtain e where "e>0" and e: "ball z0 e \ U" by (blast intro: openE [OF \open U\]) - have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" - if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c - proof - - have *: "\x1 x2 z. z \ U \ closed_segment x1 x2 \ U \ (\w. d w z) contour_integrable_on linepath x1 x2" - using hol_dw holomorphic_on_imp_continuous_on \open U\ - by (auto intro!: contour_integrable_holomorphic_simple) - have abc: "closed_segment a b \ U" "closed_segment b c \ U" "closed_segment c a \ U" - using that e segments_subset_convex_hull by fastforce+ - have eq0: "\w. w \ U \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" - apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) - apply (rule holomorphic_on_subset [OF hol_dw]) - using e abc_subset by auto - have "contour_integral \ - (\x. contour_integral (linepath a b) (\z. d z x) + - (contour_integral (linepath b c) (\z. d z x) + - contour_integral (linepath c a) (\z. d z x))) = 0" - apply (rule contour_integral_eq_0) - using abc pasz U - apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ - done - then show ?thesis - by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) - qed - show ?thesis - using e \e > 0\ - by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic - Morera_triangle continuous_on_subset [OF conthu] *) - qed - qed - qed - ultimately have [simp]: "h z = 0" for z - by (meson Liouville_weak) - have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" - by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) - then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" - by (metis mult.commute has_contour_integral_lmul) - then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" - by (simp add: field_split_simps) - moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" - using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) - show ?thesis - using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) -qed - -theorem Cauchy_integral_formula_global: - assumes S: "open S" and holf: "f holomorphic_on S" - and z: "z \ S" and vpg: "valid_path \" - and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" - and zero: "\w. w \ S \ winding_number \ w = 0" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - have "path \" using vpg by (blast intro: valid_path_imp_path) - have hols: "(\w. f w / (w - z)) holomorphic_on S - {z}" "(\w. 1 / (w - z)) holomorphic_on S - {z}" - by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ - then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" - by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete S vpg pasz) - obtain d where "d>0" - and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; - pathstart h = pathstart g \ pathfinish h = pathfinish g\ - \ path_image h \ S - {z} \ (\f. f holomorphic_on S - {z} \ contour_integral h f = contour_integral g f)" - using contour_integral_nearby_ends [OF _ \path \\ pasz] S by (simp add: open_Diff) metis - obtain p where polyp: "polynomial_function p" - and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" - using path_approx_polynomial_function [OF \path \\ \d > 0\] by blast - then have ploop: "pathfinish p = pathstart p" using loop by auto - have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast - have [simp]: "z \ path_image \" using pasz by blast - have paps: "path_image p \ S - {z}" and cint_eq: "(\f. f holomorphic_on S - {z} \ contour_integral p f = contour_integral \ f)" - using pf ps led d [OF vpg vpp] \d > 0\ by auto - have wn_eq: "winding_number p z = winding_number \ z" - using vpp paps - by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) - have "winding_number p w = winding_number \ w" if "w \ S" for w - proof - - have hol: "(\v. 1 / (v - w)) holomorphic_on S - {z}" - using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) - have "w \ path_image p" "w \ path_image \" using paps pasz that by auto - then show ?thesis - using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) - qed - then have wn0: "\w. w \ S \ winding_number p w = 0" - by (simp add: zero) - show ?thesis - using Cauchy_integral_formula_global_weak [OF S holf z polyp paps ploop wn0] hols - by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) -qed - -theorem Cauchy_theorem_global: - assumes S: "open S" and holf: "f holomorphic_on S" - and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" - and pas: "path_image \ \ S" - and zero: "\w. w \ S \ winding_number \ w = 0" - shows "(f has_contour_integral 0) \" -proof - - obtain z where "z \ S" and znot: "z \ path_image \" - proof - - have "compact (path_image \)" - using compact_valid_path_image vpg by blast - then have "path_image \ \ S" - by (metis (no_types) compact_open path_image_nonempty S) - with pas show ?thesis by (blast intro: that) - qed - then have pasz: "path_image \ \ S - {z}" using pas by blast - have hol: "(\w. (w - z) * f w) holomorphic_on S" - by (rule holomorphic_intros holf)+ - show ?thesis - using Cauchy_integral_formula_global [OF S hol \z \ S\ vpg pasz loop zero] - by (auto simp: znot elim!: has_contour_integral_eq) -qed - -corollary Cauchy_theorem_global_outside: - assumes "open S" "f holomorphic_on S" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ S" - "\w. w \ S \ w \ outside(path_image \)" - shows "(f has_contour_integral 0) \" -by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) - -lemma simply_connected_imp_winding_number_zero: - assumes "simply_connected S" "path g" - "path_image g \ S" "pathfinish g = pathstart g" "z \ S" - shows "winding_number g z = 0" -proof - - have hom: "homotopic_loops S g (linepath (pathstart g) (pathstart g))" - by (meson assms homotopic_paths_imp_homotopic_loops pathfinish_linepath simply_connected_eq_contractible_path) - then have "homotopic_paths (- {z}) g (linepath (pathstart g) (pathstart g))" - by (meson \z \ S\ homotopic_loops_imp_homotopic_paths_null homotopic_paths_subset subset_Compl_singleton) - then have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z" - by (rule winding_number_homotopic_paths) - also have "\ = 0" - using assms by (force intro: winding_number_trivial) - finally show ?thesis . -qed - -lemma Cauchy_theorem_simply_connected: - assumes "open S" "simply_connected S" "f holomorphic_on S" "valid_path g" - "path_image g \ S" "pathfinish g = pathstart g" - shows "(f has_contour_integral 0) g" -using assms -apply (simp add: simply_connected_eq_contractible_path) -apply (auto intro!: Cauchy_theorem_null_homotopic [where a = "pathstart g"] - homotopic_paths_imp_homotopic_loops) -using valid_path_imp_path by blast - -proposition\<^marker>\tag unimportant\ holomorphic_logarithm_exists: - assumes A: "convex A" "open A" - and f: "f holomorphic_on A" "\x. x \ A \ f x \ 0" - and z0: "z0 \ A" - obtains g where "g holomorphic_on A" and "\x. x \ A \ exp (g x) = f x" -proof - - note f' = holomorphic_derivI [OF f(1) A(2)] - obtain g where g: "\x. x \ A \ (g has_field_derivative deriv f x / f x) (at x)" - proof (rule holomorphic_convex_primitive' [OF A]) - show "(\x. deriv f x / f x) holomorphic_on A" - by (intro holomorphic_intros f A) - qed (auto simp: A at_within_open[of _ A]) - define h where "h = (\x. -g z0 + ln (f z0) + g x)" - from g and A have g_holo: "g holomorphic_on A" - by (auto simp: holomorphic_on_def at_within_open[of _ A] field_differentiable_def) - hence h_holo: "h holomorphic_on A" - by (auto simp: h_def intro!: holomorphic_intros) - have "\c. \x\A. f x / exp (h x) - 1 = c" - proof (rule has_field_derivative_zero_constant, goal_cases) - case (2 x) - note [simp] = at_within_open[OF _ \open A\] - from 2 and z0 and f show ?case - by (auto simp: h_def exp_diff field_simps intro!: derivative_eq_intros g f') - qed fact+ - then obtain c where c: "\x. x \ A \ f x / exp (h x) - 1 = c" - by blast - from c[OF z0] and z0 and f have "c = 0" - by (simp add: h_def) - with c have "\x. x \ A \ exp (h x) = f x" by simp - from that[OF h_holo this] show ?thesis . -qed end diff --git a/src/HOL/Analysis/Complex_Transcendental.thy b/src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy +++ b/src/HOL/Analysis/Complex_Transcendental.thy @@ -1,4056 +1,4058 @@ section \Complex Transcendental Functions\ text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ theory Complex_Transcendental imports Complex_Analysis_Basics Summation_Tests "HOL-Library.Periodic_Fun" begin subsection\Möbius transformations\ (* TODO: Figure out what to do with Möbius transformations *) definition\<^marker>\tag important\ "moebius a b c d = (\z. (a*z+b) / (c*z+d :: 'a :: field))" theorem moebius_inverse: assumes "a * d \ b * c" "c * z + d \ 0" shows "moebius d (-b) (-c) a (moebius a b c d z) = z" proof - from assms have "(-c) * moebius a b c d z + a \ 0" unfolding moebius_def by (simp add: field_simps) with assms show ?thesis unfolding moebius_def by (simp add: moebius_def divide_simps) (simp add: algebra_simps)? qed lemma moebius_inverse': assumes "a * d \ b * c" "c * z - a \ 0" shows "moebius a b c d (moebius d (-b) (-c) a z) = z" using assms moebius_inverse[of d a "-b" "-c" z] by (auto simp: algebra_simps) lemma cmod_add_real_less: assumes "Im z \ 0" "r\0" shows "cmod (z + r) < cmod z + \r\" proof (cases z) case (Complex x y) have "r * x / \r\ < sqrt (x*x + y*y)" apply (rule real_less_rsqrt) using assms apply (simp add: Complex power2_eq_square) using not_real_square_gt_zero by blast then show ?thesis using assms Complex apply (simp add: cmod_def) apply (rule power2_less_imp_less, auto) apply (simp add: power2_eq_square field_simps) done qed lemma cmod_diff_real_less: "Im z \ 0 \ x\0 \ cmod (z - x) < cmod z + \x\" using cmod_add_real_less [of z "-x"] by simp lemma cmod_square_less_1_plus: assumes "Im z = 0 \ \Re z\ < 1" shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)" proof (cases "Im z = 0 \ Re z = 0") case True with assms abs_square_less_1 show ?thesis by (force simp add: Re_power2 Im_power2 cmod_def) next case False with cmod_diff_real_less [of "1 - z\<^sup>2" "1"] show ?thesis by (simp add: norm_power Im_power2) qed subsection\<^marker>\tag unimportant\\The Exponential Function\ lemma norm_exp_i_times [simp]: "norm (exp(\ * of_real y)) = 1" by simp lemma norm_exp_imaginary: "norm(exp z) = 1 \ Re z = 0" by simp lemma field_differentiable_within_exp: "exp field_differentiable (at z within s)" using DERIV_exp field_differentiable_at_within field_differentiable_def by blast lemma continuous_within_exp: fixes z::"'a::{real_normed_field,banach}" shows "continuous (at z within s) exp" by (simp add: continuous_at_imp_continuous_within) lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s" by (simp add: field_differentiable_within_exp holomorphic_on_def) lemma holomorphic_on_exp' [holomorphic_intros]: "f holomorphic_on s \ (\x. exp (f x)) holomorphic_on s" using holomorphic_on_compose[OF _ holomorphic_on_exp] by (simp add: o_def) subsection\Euler and de Moivre formulas\ text\The sine series times \<^term>\i\\ lemma sin_i_eq: "(\n. (\ * sin_coeff n) * z^n) sums (\ * sin z)" proof - have "(\n. \ * sin_coeff n *\<^sub>R z^n) sums (\ * sin z)" using sin_converges sums_mult by blast then show ?thesis by (simp add: scaleR_conv_of_real field_simps) qed theorem exp_Euler: "exp(\ * z) = cos(z) + \ * sin(z)" proof - have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) = (\n. (\ * z) ^ n /\<^sub>R (fact n))" proof fix n show "(cos_coeff n + \ * sin_coeff n) * z^n = (\ * z) ^ n /\<^sub>R (fact n)" by (auto simp: cos_coeff_def sin_coeff_def scaleR_conv_of_real field_simps elim!: evenE oddE) qed also have "... sums (exp (\ * z))" by (rule exp_converges) finally have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) sums (exp (\ * z))" . moreover have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) sums (cos z + \ * sin z)" using sums_add [OF cos_converges [of z] sin_i_eq [of z]] by (simp add: field_simps scaleR_conv_of_real) ultimately show ?thesis using sums_unique2 by blast qed corollary\<^marker>\tag unimportant\ exp_minus_Euler: "exp(-(\ * z)) = cos(z) - \ * sin(z)" using exp_Euler [of "-z"] by simp lemma sin_exp_eq: "sin z = (exp(\ * z) - exp(-(\ * z))) / (2*\)" by (simp add: exp_Euler exp_minus_Euler) lemma sin_exp_eq': "sin z = \ * (exp(-(\ * z)) - exp(\ * z)) / 2" by (simp add: exp_Euler exp_minus_Euler) lemma cos_exp_eq: "cos z = (exp(\ * z) + exp(-(\ * z))) / 2" by (simp add: exp_Euler exp_minus_Euler) theorem Euler: "exp(z) = of_real(exp(Re z)) * (of_real(cos(Im z)) + \ * of_real(sin(Im z)))" by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real Complex_eq) lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2" by (simp add: sin_exp_eq field_simps Re_divide Im_exp) lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2" by (simp add: sin_exp_eq field_simps Im_divide Re_exp) lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2" by (simp add: cos_exp_eq field_simps Re_divide Re_exp) lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2" by (simp add: cos_exp_eq field_simps Im_divide Im_exp) lemma Re_sin_pos: "0 < Re z \ Re z < pi \ Re (sin z) > 0" by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero) lemma Im_sin_nonneg: "Re z = 0 \ 0 \ Im z \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) lemma Im_sin_nonneg2: "Re z = pi \ Im z \ 0 \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) subsection\<^marker>\tag unimportant\\Relationships between real and complex trigonometric and hyperbolic functions\ lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x" by (simp add: sin_of_real) lemma real_cos_eq [simp]: "Re(cos(of_real x)) = cos x" by (simp add: cos_of_real) lemma DeMoivre: "(cos z + \ * sin z) ^ n = cos(n * z) + \ * sin(n * z)" by (metis exp_Euler [symmetric] exp_of_nat_mult mult.left_commute) lemma exp_cnj: "cnj (exp z) = exp (cnj z)" proof - have "(\n. cnj (z ^ n /\<^sub>R (fact n))) = (\n. (cnj z)^n /\<^sub>R (fact n))" by auto also have "... sums (exp (cnj z))" by (rule exp_converges) finally have "(\n. cnj (z ^ n /\<^sub>R (fact n))) sums (exp (cnj z))" . moreover have "(\n. cnj (z ^ n /\<^sub>R (fact n))) sums (cnj (exp z))" by (metis exp_converges sums_cnj) ultimately show ?thesis using sums_unique2 by blast qed lemma cnj_sin: "cnj(sin z) = sin(cnj z)" by (simp add: sin_exp_eq exp_cnj field_simps) lemma cnj_cos: "cnj(cos z) = cos(cnj z)" by (simp add: cos_exp_eq exp_cnj field_simps) lemma field_differentiable_at_sin: "sin field_differentiable at z" using DERIV_sin field_differentiable_def by blast lemma field_differentiable_within_sin: "sin field_differentiable (at z within S)" by (simp add: field_differentiable_at_sin field_differentiable_at_within) lemma field_differentiable_at_cos: "cos field_differentiable at z" using DERIV_cos field_differentiable_def by blast lemma field_differentiable_within_cos: "cos field_differentiable (at z within S)" by (simp add: field_differentiable_at_cos field_differentiable_at_within) lemma holomorphic_on_sin: "sin holomorphic_on S" by (simp add: field_differentiable_within_sin holomorphic_on_def) lemma holomorphic_on_cos: "cos holomorphic_on S" by (simp add: field_differentiable_within_cos holomorphic_on_def) lemma holomorphic_on_sin' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. sin (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_sin] by (simp add: o_def) lemma holomorphic_on_cos' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. cos (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def) subsection\<^marker>\tag unimportant\\More on the Polar Representation of Complex Numbers\ lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)" by (simp add: Complex_eq exp_add exp_Euler exp_of_real sin_of_real cos_of_real) lemma exp_eq_1: "exp z = 1 \ Re(z) = 0 \ (\n::int. Im(z) = of_int (2 * n) * pi)" (is "?lhs = ?rhs") proof assume "exp z = 1" then have "Re z = 0" by (metis exp_eq_one_iff norm_exp_eq_Re norm_one) with \?lhs\ show ?rhs by (metis Re_exp complex_Re_of_int cos_one_2pi_int exp_zero mult.commute mult_numeral_1 numeral_One of_int_mult of_int_numeral) next assume ?rhs then show ?lhs using Im_exp Re_exp complex_eq_iff by (simp add: cos_one_2pi_int cos_one_sin_zero mult.commute) qed lemma exp_eq: "exp w = exp z \ (\n::int. w = z + (of_int (2 * n) * pi) * \)" (is "?lhs = ?rhs") proof - have "exp w = exp z \ exp (w-z) = 1" by (simp add: exp_diff) also have "... \ (Re w = Re z \ (\n::int. Im w - Im z = of_int (2 * n) * pi))" by (simp add: exp_eq_1) also have "... \ ?rhs" by (auto simp: algebra_simps intro!: complex_eqI) finally show ?thesis . qed lemma exp_complex_eqI: "\Im w - Im z\ < 2*pi \ exp w = exp z \ w = z" by (auto simp: exp_eq abs_mult) lemma exp_integer_2pi: assumes "n \ \" shows "exp((2 * n * pi) * \) = 1" proof - have "exp((2 * n * pi) * \) = exp 0" using assms unfolding Ints_def exp_eq by auto also have "... = 1" by simp finally show ?thesis . qed lemma exp_plus_2pin [simp]: "exp (z + \ * (of_int n * (of_real pi * 2))) = exp z" by (simp add: exp_eq) lemma exp_integer_2pi_plus1: assumes "n \ \" shows "exp(((2 * n + 1) * pi) * \) = - 1" proof - from assms obtain n' where [simp]: "n = of_int n'" by (auto simp: Ints_def) have "exp(((2 * n + 1) * pi) * \) = exp (pi * \)" using assms by (subst exp_eq) (auto intro!: exI[of _ n'] simp: algebra_simps) also have "... = - 1" by simp finally show ?thesis . qed lemma inj_on_exp_pi: fixes z::complex shows "inj_on exp (ball z pi)" proof (clarsimp simp: inj_on_def exp_eq) fix y n assume "dist z (y + 2 * of_int n * of_real pi * \) < pi" "dist z y < pi" then have "dist y (y + 2 * of_int n * of_real pi * \) < pi+pi" using dist_commute_lessI dist_triangle_less_add by blast then have "norm (2 * of_int n * of_real pi * \) < 2*pi" by (simp add: dist_norm) then show "n = 0" by (auto simp: norm_mult) qed lemma cmod_add_squared: fixes r1 r2::real assumes "r1 \ 0" "r2 \ 0" shows "(cmod (r1 * exp (\ * \1) + r2 * exp (\ * \2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 + 2 * r1 * r2 * cos (\1 - \2)" (is "(cmod (?z1 + ?z2))\<^sup>2 = ?rhs") proof - have "(cmod (?z1 + ?z2))\<^sup>2 = (?z1 + ?z2) * cnj (?z1 + ?z2)" by (rule complex_norm_square) also have "\ = (?z1 * cnj ?z1 + ?z2 * cnj ?z2) + (?z1 * cnj ?z2 + cnj ?z1 * ?z2)" by (simp add: algebra_simps) also have "\ = (norm ?z1)\<^sup>2 + (norm ?z2)\<^sup>2 + 2 * Re (?z1 * cnj ?z2)" unfolding complex_norm_square [symmetric] cnj_add_mult_eq_Re by simp also have "\ = ?rhs" by (simp add: norm_mult) (simp add: exp_Euler complex_is_Real_iff [THEN iffD1] cos_diff algebra_simps) finally show ?thesis using of_real_eq_iff by blast qed lemma cmod_diff_squared: fixes r1 r2::real assumes "r1 \ 0" "r2 \ 0" shows "(cmod (r1 * exp (\ * \1) - r2 * exp (\ * \2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\1 - \2)" (is "(cmod (?z1 - ?z2))\<^sup>2 = ?rhs") proof - have "exp (\ * (\2 + pi)) = - exp (\ * \2)" by (simp add: exp_Euler cos_plus_pi sin_plus_pi) then have "(cmod (?z1 - ?z2))\<^sup>2 = cmod (?z1 + r2 * exp (\ * (\2 + pi))) ^2" by simp also have "\ = r1\<^sup>2 + r2\<^sup>2 + 2*r1*r2*cos (\1 - (\2 + pi))" using assms cmod_add_squared by blast also have "\ = ?rhs" by (simp add: add.commute diff_add_eq_diff_diff_swap) finally show ?thesis . qed lemma polar_convergence: fixes R::real assumes "\j. r j > 0" "R > 0" shows "((\j. r j * exp (\ * \ j)) \ (R * exp (\ * \))) \ (r \ R) \ (\k. (\j. \ j - of_int (k j) * (2 * pi)) \ \)" (is "(?z \ ?Z) = ?rhs") proof assume L: "?z \ ?Z" have rR: "r \ R" using tendsto_norm [OF L] assms by (auto simp: norm_mult abs_of_pos) moreover obtain k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" proof - have "cos (\ j - \) = ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)" for j apply (subst cmod_diff_squared) using assms by (auto simp: field_split_simps less_le) moreover have "(\j. ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)) \ ((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R))" by (intro L rR tendsto_intros) (use \R > 0\ in force) moreover have "((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R)) = 1" using \R > 0\ by (simp add: power2_eq_square field_split_simps) ultimately have "(\j. cos (\ j - \)) \ 1" by auto then show ?thesis using that cos_diff_limit_1 by blast qed ultimately show ?rhs by metis next assume R: ?rhs show "?z \ ?Z" proof (rule tendsto_mult) show "(\x. complex_of_real (r x)) \ of_real R" using R by (auto simp: tendsto_of_real_iff) obtain k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" using R by metis then have "(\j. complex_of_real (\ j - of_int (k j) * (2 * pi))) \ of_real \" using tendsto_of_real_iff by force then have "(\j. exp (\ * of_real (\ j - of_int (k j) * (2 * pi)))) \ exp (\ * \)" using tendsto_mult [OF tendsto_const] isCont_exp isCont_tendsto_compose by blast moreover have "exp (\ * of_real (\ j - of_int (k j) * (2 * pi))) = exp (\ * \ j)" for j unfolding exp_eq by (rule_tac x="- k j" in exI) (auto simp: algebra_simps) ultimately show "(\j. exp (\ * \ j)) \ exp (\ * \)" by auto qed qed lemma sin_cos_eq_iff: "sin y = sin x \ cos y = cos x \ (\n::int. y = x + 2 * pi * n)" proof - { assume "sin y = sin x" "cos y = cos x" then have "cos (y-x) = 1" using cos_add [of y "-x"] by simp then have "\n::int. y-x = 2 * pi * n" using cos_one_2pi_int by auto } then show ?thesis apply (auto simp: sin_add cos_add) apply (metis add.commute diff_add_cancel) done qed lemma exp_i_ne_1: assumes "0 < x" "x < 2*pi" shows "exp(\ * of_real x) \ 1" proof assume "exp (\ * of_real x) = 1" then have "exp (\ * of_real x) = exp 0" by simp then obtain n where "\ * of_real x = (of_int (2 * n) * pi) * \" by (simp only: Ints_def exp_eq) auto then have "of_real x = (of_int (2 * n) * pi)" by (metis complex_i_not_zero mult.commute mult_cancel_left of_real_eq_iff real_scaleR_def scaleR_conv_of_real) then have "x = (of_int (2 * n) * pi)" by simp then show False using assms by (cases n) (auto simp: zero_less_mult_iff mult_less_0_iff) qed lemma sin_eq_0: fixes z::complex shows "sin z = 0 \ (\n::int. z = of_real(n * pi))" by (simp add: sin_exp_eq exp_eq) lemma cos_eq_0: fixes z::complex shows "cos z = 0 \ (\n::int. z = of_real(n * pi) + of_real pi/2)" using sin_eq_0 [of "z - of_real pi/2"] by (simp add: sin_diff algebra_simps) lemma cos_eq_1: fixes z::complex shows "cos z = 1 \ (\n::int. z = of_real(2 * n * pi))" proof - have "cos z = cos (2*(z/2))" by simp also have "... = 1 - 2 * sin (z/2) ^ 2" by (simp only: cos_double_sin) finally have [simp]: "cos z = 1 \ sin (z/2) = 0" by simp show ?thesis by (auto simp: sin_eq_0) qed lemma csin_eq_1: fixes z::complex shows "sin z = 1 \ (\n::int. z = of_real(2 * n * pi) + of_real pi/2)" using cos_eq_1 [of "z - of_real pi/2"] by (simp add: cos_diff algebra_simps) lemma csin_eq_minus1: fixes z::complex shows "sin z = -1 \ (\n::int. z = of_real(2 * n * pi) + 3/2*pi)" (is "_ = ?rhs") proof - have "sin z = -1 \ sin (-z) = 1" by (simp add: equation_minus_iff) also have "... \ (\n::int. -z = of_real(2 * n * pi) + of_real pi/2)" by (simp only: csin_eq_1) also have "... \ (\n::int. z = - of_real(2 * n * pi) - of_real pi/2)" apply (rule iff_exI) by (metis (no_types) is_num_normalize(8) minus_minus of_real_def real_vector.scale_minus_left uminus_add_conv_diff) also have "... = ?rhs" apply safe apply (rule_tac [2] x="-(x+1)" in exI) apply (rule_tac x="-(x+1)" in exI) apply (simp_all add: algebra_simps) done finally show ?thesis . qed lemma ccos_eq_minus1: fixes z::complex shows "cos z = -1 \ (\n::int. z = of_real(2 * n * pi) + pi)" using csin_eq_1 [of "z - of_real pi/2"] by (simp add: sin_diff algebra_simps equation_minus_iff) lemma sin_eq_1: "sin x = 1 \ (\n::int. x = (2 * n + 1 / 2) * pi)" (is "_ = ?rhs") proof - have "sin x = 1 \ sin (complex_of_real x) = 1" by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + of_real pi/2)" by (simp only: csin_eq_1) also have "... \ (\n::int. x = of_real(2 * n * pi) + of_real pi/2)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma sin_eq_minus1: "sin x = -1 \ (\n::int. x = (2*n + 3/2) * pi)" (is "_ = ?rhs") proof - have "sin x = -1 \ sin (complex_of_real x) = -1" by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)" by (simp only: csin_eq_minus1) also have "... \ (\n::int. x = of_real(2 * n * pi) + 3/2*pi)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma cos_eq_minus1: "cos x = -1 \ (\n::int. x = (2*n + 1) * pi)" (is "_ = ?rhs") proof - have "cos x = -1 \ cos (complex_of_real x) = -1" by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + pi)" by (simp only: ccos_eq_minus1) also have "... \ (\n::int. x = of_real(2 * n * pi) + pi)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma dist_exp_i_1: "norm(exp(\ * of_real t) - 1) = 2 * \sin(t / 2)\" apply (simp add: exp_Euler cmod_def power2_diff sin_of_real cos_of_real algebra_simps) using cos_double_sin [of "t/2"] apply (simp add: real_sqrt_mult) done lemma complex_sin_eq: fixes w :: complex shows "sin w = sin z \ (\n \ \. w = z + of_real(2*n*pi) \ w = -z + of_real((2*n + 1)*pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "sin w - sin z = 0" by (auto simp: algebra_simps) then have "sin ((w - z) / 2)*cos ((w + z) / 2) = 0" by (auto simp: sin_diff_sin) then consider "sin ((w - z) / 2) = 0" | "cos ((w + z) / 2) = 0" using mult_eq_0_iff by blast then show ?rhs proof cases case 1 then show ?thesis by (simp add: sin_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq) next case 2 then show ?thesis by (simp add: cos_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq) qed next assume ?rhs then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \ w = -z + of_real ((2* of_int n + 1)*pi)" using Ints_cases by blast then show ?lhs using Periodic_Fun.sin.plus_of_int [of z n] apply (auto simp: algebra_simps) by (metis (no_types, hide_lams) add_diff_cancel_left add_diff_cancel_left' add_minus_cancel mult.commute sin.plus_of_int sin_minus sin_plus_pi) qed lemma complex_cos_eq: fixes w :: complex shows "cos w = cos z \ (\n \ \. w = z + of_real(2*n*pi) \ w = -z + of_real(2*n*pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "cos w - cos z = 0" by (auto simp: algebra_simps) then have "sin ((w + z) / 2) * sin ((z - w) / 2) = 0" by (auto simp: cos_diff_cos) then consider "sin ((w + z) / 2) = 0" | "sin ((z - w) / 2) = 0" using mult_eq_0_iff by blast then show ?rhs proof cases case 1 then show ?thesis apply (simp add: sin_eq_0 algebra_simps) by (metis Ints_of_int of_real_of_int_eq) next case 2 then show ?thesis apply (clarsimp simp: sin_eq_0 algebra_simps) by (metis Ints_of_int add_minus_cancel distrib_right mult_of_int_commute mult_zero_right of_int_0 of_int_add of_real_of_int_eq) qed next assume ?rhs then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \ w = -z + of_real(2*n*pi)" using Ints_cases by (metis of_int_mult of_int_numeral) then show ?lhs using Periodic_Fun.cos.plus_of_int [of z n] apply (simp add: algebra_simps) by (metis cos.plus_of_int cos_minus minus_add_cancel mult.commute) qed lemma sin_eq: "sin x = sin y \ (\n \ \. x = y + 2*n*pi \ x = -y + (2*n + 1)*pi)" using complex_sin_eq [of x y] by (simp only: sin_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff) lemma cos_eq: "cos x = cos y \ (\n \ \. x = y + 2*n*pi \ x = -y + 2*n*pi)" using complex_cos_eq [of x y] by (simp only: cos_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff) lemma sinh_complex: fixes z :: complex shows "(exp z - inverse (exp z)) / 2 = -\ * sin(\ * z)" by (simp add: sin_exp_eq field_split_simps exp_minus) lemma sin_i_times: fixes z :: complex shows "sin(\ * z) = \ * ((exp z - inverse (exp z)) / 2)" using sinh_complex by auto lemma sinh_real: fixes x :: real shows "of_real((exp x - inverse (exp x)) / 2) = -\ * sin(\ * of_real x)" by (simp add: exp_of_real sin_i_times) lemma cosh_complex: fixes z :: complex shows "(exp z + inverse (exp z)) / 2 = cos(\ * z)" by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real) lemma cosh_real: fixes x :: real shows "of_real((exp x + inverse (exp x)) / 2) = cos(\ * of_real x)" by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real) lemmas cos_i_times = cosh_complex [symmetric] lemma norm_cos_squared: "norm(cos z) ^ 2 = cos(Re z) ^ 2 + (exp(Im z) - inverse(exp(Im z))) ^ 2 / 4" apply (cases z) apply (simp add: cos_add cmod_power2 cos_of_real sin_of_real Complex_eq) apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) apply (simp only: left_diff_distrib [symmetric] power_mult_distrib sin_squared_eq) apply (simp add: power2_eq_square algebra_simps field_split_simps) done lemma norm_sin_squared: "norm(sin z) ^ 2 = (exp(2 * Im z) + inverse(exp(2 * Im z)) - 2 * cos(2 * Re z)) / 4" apply (cases z) apply (simp add: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double Complex_eq) apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) apply (simp only: left_diff_distrib [symmetric] power_mult_distrib cos_squared_eq) apply (simp add: power2_eq_square algebra_simps field_split_simps) done lemma exp_uminus_Im: "exp (- Im z) \ exp (cmod z)" using abs_Im_le_cmod linear order_trans by fastforce lemma norm_cos_le: fixes z::complex shows "norm(cos z) \ exp(norm z)" proof - have "Im z \ cmod z" using abs_Im_le_cmod abs_le_D1 by auto with exp_uminus_Im show ?thesis apply (simp add: cos_exp_eq norm_divide) apply (rule order_trans [OF norm_triangle_ineq], simp) apply (metis add_mono exp_le_cancel_iff mult_2_right) done qed lemma norm_cos_plus1_le: fixes z::complex shows "norm(1 + cos z) \ 2 * exp(norm z)" proof - have mono: "\u w z::real. (1 \ w | 1 \ z) \ (w \ u & z \ u) \ 2 + w + z \ 4 * u" by arith have *: "Im z \ cmod z" using abs_Im_le_cmod abs_le_D1 by auto have triangle3: "\x y z. norm(x + y + z) \ norm(x) + norm(y) + norm(z)" by (simp add: norm_add_rule_thm) have "norm(1 + cos z) = cmod (1 + (exp (\ * z) + exp (- (\ * z))) / 2)" by (simp add: cos_exp_eq) also have "... = cmod ((2 + exp (\ * z) + exp (- (\ * z))) / 2)" by (simp add: field_simps) also have "... = cmod (2 + exp (\ * z) + exp (- (\ * z))) / 2" by (simp add: norm_divide) finally show ?thesis by (metis exp_eq_one_iff exp_le_cancel_iff mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono) qed lemma sinh_conv_sin: "sinh z = -\ * sin (\*z)" by (simp add: sinh_field_def sin_i_times exp_minus) lemma cosh_conv_cos: "cosh z = cos (\*z)" by (simp add: cosh_field_def cos_i_times exp_minus) lemma tanh_conv_tan: "tanh z = -\ * tan (\*z)" by (simp add: tanh_def sinh_conv_sin cosh_conv_cos tan_def) lemma sin_conv_sinh: "sin z = -\ * sinh (\*z)" by (simp add: sinh_conv_sin) lemma cos_conv_cosh: "cos z = cosh (\*z)" by (simp add: cosh_conv_cos) lemma tan_conv_tanh: "tan z = -\ * tanh (\*z)" by (simp add: tan_def sin_conv_sinh cos_conv_cosh tanh_def) lemma sinh_complex_eq_iff: "sinh (z :: complex) = sinh w \ (\n\\. z = w - 2 * \ * of_real n * of_real pi \ z = -(2 * complex_of_real n + 1) * \ * complex_of_real pi - w)" (is "_ = ?rhs") proof - have "sinh z = sinh w \ sin (\ * z) = sin (\ * w)" by (simp add: sinh_conv_sin) also have "\ \ ?rhs" by (subst complex_sin_eq) (force simp: field_simps complex_eq_iff) finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Taylor series for complex exponential, sine and cosine\ declare power_Suc [simp del] lemma Taylor_exp_field: fixes z::"'a::{banach,real_normed_field}" shows "norm (exp z - (\i\n. z ^ i / fact i)) \ exp (norm z) * (norm z ^ Suc n) / fact n" proof (rule field_Taylor[of _ n "\k. exp" "exp (norm z)" 0 z, simplified]) show "convex (closed_segment 0 z)" by (rule convex_closed_segment [of 0 z]) next fix k x assume "x \ closed_segment 0 z" "k \ n" show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" using DERIV_exp DERIV_subset by blast next fix x assume x: "x \ closed_segment 0 z" have "norm (exp x) \ exp (norm x)" by (rule norm_exp) also have "norm x \ norm z" using x by (auto simp: closed_segment_def intro!: mult_left_le_one_le) finally show "norm (exp x) \ exp (norm z)" by simp qed auto lemma Taylor_exp: "norm(exp z - (\k\n. z ^ k / (fact k))) \ exp\Re z\ * (norm z) ^ (Suc n) / (fact n)" proof (rule complex_Taylor [of _ n "\k. exp" "exp\Re z\" 0 z, simplified]) show "convex (closed_segment 0 z)" by (rule convex_closed_segment [of 0 z]) next fix k x assume "x \ closed_segment 0 z" "k \ n" show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" using DERIV_exp DERIV_subset by blast next fix x assume "x \ closed_segment 0 z" then show "Re x \ \Re z\" apply (clarsimp simp: closed_segment_def scaleR_conv_of_real) by (meson abs_ge_self abs_ge_zero linear mult_left_le_one_le mult_nonneg_nonpos order_trans) qed auto lemma assumes "0 \ u" "u \ 1" shows cmod_sin_le_exp: "cmod (sin (u *\<^sub>R z)) \ exp \Im z\" and cmod_cos_le_exp: "cmod (cos (u *\<^sub>R z)) \ exp \Im z\" proof - have mono: "\u w z::real. w \ u \ z \ u \ (w + z)/2 \ u" by simp have *: "(cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2 \ exp \Im z\" proof (rule mono) show "cmod (exp (\ * (u * z))) \ exp \Im z\" apply (simp add: abs_if mult_left_le_one_le assms not_less) by (meson assms(1) mult_nonneg_nonneg neg_le_0_iff_le order_trans) show "cmod (exp (- (\ * (u * z)))) \ exp \Im z\" apply (simp add: abs_if mult_left_le_one_le assms) by (meson \0 \ u\ less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans) qed have "cmod (sin (u *\<^sub>R z)) = cmod (exp (\ * (u * z)) - exp (- (\ * (u * z)))) / 2" by (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide) also have "... \ (cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2" by (intro divide_right_mono norm_triangle_ineq4) simp also have "... \ exp \Im z\" by (rule *) finally show "cmod (sin (u *\<^sub>R z)) \ exp \Im z\" . have "cmod (cos (u *\<^sub>R z)) = cmod (exp (\ * (u * z)) + exp (- (\ * (u * z)))) / 2" by (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide) also have "... \ (cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2" by (intro divide_right_mono norm_triangle_ineq) simp also have "... \ exp \Im z\" by (rule *) finally show "cmod (cos (u *\<^sub>R z)) \ exp \Im z\" . qed lemma Taylor_sin: "norm(sin z - (\k\n. complex_of_real (sin_coeff k) * z ^ k)) \ exp\Im z\ * (norm z) ^ (Suc n) / (fact n)" proof - have mono: "\u w z::real. w \ u \ z \ u \ w + z \ u*2" by arith have *: "cmod (sin z - (\i\n. (-1) ^ (i div 2) * (if even i then sin 0 else cos 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" proof (rule complex_Taylor [of "closed_segment 0 z" n "\k x. (-1)^(k div 2) * (if even k then sin x else cos x)" "exp\Im z\" 0 z, simplified]) fix k x show "((\x. (- 1) ^ (k div 2) * (if even k then sin x else cos x)) has_field_derivative (- 1) ^ (Suc k div 2) * (if odd k then sin x else cos x)) (at x within closed_segment 0 z)" apply (auto simp: power_Suc) apply (intro derivative_eq_intros | simp)+ done next fix x assume "x \ closed_segment 0 z" then show "cmod ((- 1) ^ (Suc n div 2) * (if odd n then sin x else cos x)) \ exp \Im z\" by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) qed have **: "\k. complex_of_real (sin_coeff k) * z ^ k = (-1)^(k div 2) * (if even k then sin 0 else cos 0) * z^k / of_nat (fact k)" by (auto simp: sin_coeff_def elim!: oddE) show ?thesis apply (rule order_trans [OF _ *]) apply (simp add: **) done qed lemma Taylor_cos: "norm(cos z - (\k\n. complex_of_real (cos_coeff k) * z ^ k)) \ exp\Im z\ * (norm z) ^ Suc n / (fact n)" proof - have mono: "\u w z::real. w \ u \ z \ u \ w + z \ u*2" by arith have *: "cmod (cos z - (\i\n. (-1) ^ (Suc i div 2) * (if even i then cos 0 else sin 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" proof (rule complex_Taylor [of "closed_segment 0 z" n "\k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)" "exp\Im z\" 0 z, simplified]) fix k x assume "x \ closed_segment 0 z" "k \ n" show "((\x. (- 1) ^ (Suc k div 2) * (if even k then cos x else sin x)) has_field_derivative (- 1) ^ Suc (k div 2) * (if odd k then cos x else sin x)) (at x within closed_segment 0 z)" apply (auto simp: power_Suc) apply (intro derivative_eq_intros | simp)+ done next fix x assume "x \ closed_segment 0 z" then show "cmod ((- 1) ^ Suc (n div 2) * (if odd n then cos x else sin x)) \ exp \Im z\" by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) qed have **: "\k. complex_of_real (cos_coeff k) * z ^ k = (-1)^(Suc k div 2) * (if even k then cos 0 else sin 0) * z^k / of_nat (fact k)" by (auto simp: cos_coeff_def elim!: evenE) show ?thesis apply (rule order_trans [OF _ *]) apply (simp add: **) done qed declare power_Suc [simp] text\32-bit Approximation to e\ lemma e_approx_32: "\exp(1) - 5837465777 / 2147483648\ \ (inverse(2 ^ 32)::real)" using Taylor_exp [of 1 14] exp_le apply (simp add: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral) apply (simp only: pos_le_divide_eq [symmetric]) done lemma e_less_272: "exp 1 < (272/100::real)" using e_approx_32 by (simp add: abs_if split: if_split_asm) lemma ln_272_gt_1: "ln (272/100) > (1::real)" by (metis e_less_272 exp_less_cancel_iff exp_ln_iff less_trans ln_exp) text\Apparently redundant. But many arguments involve integers.\ lemma ln3_gt_1: "ln 3 > (1::real)" by (simp add: less_trans [OF ln_272_gt_1]) subsection\The argument of a complex number (HOL Light version)\ definition\<^marker>\tag important\ is_Arg :: "[complex,real] \ bool" where "is_Arg z r \ z = of_real(norm z) * exp(\ * of_real r)" definition\<^marker>\tag important\ Arg2pi :: "complex \ real" where "Arg2pi z \ if z = 0 then 0 else THE t. 0 \ t \ t < 2*pi \ is_Arg z t" lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \ is_Arg z r" by (simp add: algebra_simps is_Arg_def) lemma is_Arg_eqI: assumes r: "is_Arg z r" and s: "is_Arg z s" and rs: "abs (r-s) < 2*pi" and "z \ 0" shows "r = s" proof - have zr: "z = (cmod z) * exp (\ * r)" and zs: "z = (cmod z) * exp (\ * s)" using r s by (auto simp: is_Arg_def) with \z \ 0\ have eq: "exp (\ * r) = exp (\ * s)" by (metis mult_eq_0_iff mult_left_cancel) have "\ * r = \ * s" by (rule exp_complex_eqI) (use rs in \auto simp: eq exp_complex_eqI\) then show ?thesis by simp qed text\This function returns the angle of a complex number from its representation in polar coordinates. Due to periodicity, its range is arbitrary. \<^term>\Arg2pi\ follows HOL Light in adopting the interval \[0,2\)\. But we have the same periodicity issue with logarithms, and it is usual to adopt the same interval for the complex logarithm and argument functions. Further on down, we shall define both functions for the interval \(-\,\]\. The present version is provided for compatibility.\ lemma Arg2pi_0 [simp]: "Arg2pi(0) = 0" by (simp add: Arg2pi_def) lemma Arg2pi_unique_lemma: assumes z: "is_Arg z t" and z': "is_Arg z t'" and t: "0 \ t" "t < 2*pi" and t': "0 \ t'" "t' < 2*pi" and nz: "z \ 0" shows "t' = t" proof - have [dest]: "\x y z::real. x\0 \ x+y < z \ y * of_real t') = of_real (cmod z) * exp (\ * of_real t)" by (metis z z' is_Arg_def) then have "exp (\ * of_real t') = exp (\ * of_real t)" by (metis nz mult_left_cancel mult_zero_left z is_Arg_def) then have "sin t' = sin t \ cos t' = cos t" apply (simp add: exp_Euler sin_of_real cos_of_real) by (metis Complex_eq complex.sel) then obtain n::int where n: "t' = t + 2 * n * pi" by (auto simp: sin_cos_eq_iff) then have "n=0" by (cases n) (use t t' in \auto simp: mult_less_0_iff algebra_simps\) then show "t' = t" by (simp add: n) qed lemma Arg2pi: "0 \ Arg2pi z \ Arg2pi z < 2*pi \ is_Arg z (Arg2pi z)" proof (cases "z=0") case True then show ?thesis by (simp add: Arg2pi_def is_Arg_def) next case False obtain t where t: "0 \ t" "t < 2*pi" and ReIm: "Re z / cmod z = cos t" "Im z / cmod z = sin t" using sincos_total_2pi [OF complex_unit_circle [OF False]] by blast have z: "is_Arg z t" unfolding is_Arg_def apply (rule complex_eqI) using t False ReIm apply (auto simp: exp_Euler sin_of_real cos_of_real field_split_simps) done show ?thesis apply (simp add: Arg2pi_def False) apply (rule theI [where a=t]) using t z False apply (auto intro: Arg2pi_unique_lemma) done qed corollary\<^marker>\tag unimportant\ shows Arg2pi_ge_0: "0 \ Arg2pi z" and Arg2pi_lt_2pi: "Arg2pi z < 2*pi" and Arg2pi_eq: "z = of_real(norm z) * exp(\ * of_real(Arg2pi z))" using Arg2pi is_Arg_def by auto lemma complex_norm_eq_1_exp: "norm z = 1 \ exp(\ * of_real (Arg2pi z)) = z" by (metis Arg2pi_eq cis_conv_exp mult.left_neutral norm_cis of_real_1) lemma Arg2pi_unique: "\of_real r * exp(\ * of_real a) = z; 0 < r; 0 \ a; a < 2*pi\ \ Arg2pi z = a" by (rule Arg2pi_unique_lemma [unfolded is_Arg_def, OF _ Arg2pi_eq]) (use Arg2pi [of z] in \auto simp: norm_mult\) lemma Arg2pi_minus: "z \ 0 \ Arg2pi (-z) = (if Arg2pi z < pi then Arg2pi z + pi else Arg2pi z - pi)" apply (rule Arg2pi_unique [of "norm z"]) apply (rule complex_eqI) using Arg2pi_ge_0 [of z] Arg2pi_eq [of z] Arg2pi_lt_2pi [of z] apply (auto simp: Re_exp Im_exp cos_diff sin_diff cis_conv_exp [symmetric]) apply (metis Re_rcis Im_rcis rcis_def)+ done lemma Arg2pi_times_of_real [simp]: assumes "0 < r" shows "Arg2pi (of_real r * z) = Arg2pi z" proof (cases "z=0") case False show ?thesis by (rule Arg2pi_unique [of "r * norm z"]) (use Arg2pi False assms is_Arg_def in auto) qed auto lemma Arg2pi_times_of_real2 [simp]: "0 < r \ Arg2pi (z * of_real r) = Arg2pi z" by (metis Arg2pi_times_of_real mult.commute) lemma Arg2pi_divide_of_real [simp]: "0 < r \ Arg2pi (z / of_real r) = Arg2pi z" by (metis Arg2pi_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) lemma Arg2pi_le_pi: "Arg2pi z \ pi \ 0 \ Im z" proof (cases "z=0") case False have "0 \ Im z \ 0 \ Im (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... = (0 \ Im (exp (\ * complex_of_real (Arg2pi z))))" using False by (simp add: zero_le_mult_iff) also have "... \ Arg2pi z \ pi" by (simp add: Im_exp) (metis Arg2pi_ge_0 Arg2pi_lt_2pi sin_lt_zero sin_ge_zero not_le) finally show ?thesis by blast qed auto lemma Arg2pi_lt_pi: "0 < Arg2pi z \ Arg2pi z < pi \ 0 < Im z" proof (cases "z=0") case False have "0 < Im z \ 0 < Im (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... = (0 < Im (exp (\ * complex_of_real (Arg2pi z))))" using False by (simp add: zero_less_mult_iff) also have "... \ 0 < Arg2pi z \ Arg2pi z < pi" using Arg2pi_ge_0 Arg2pi_lt_2pi sin_le_zero sin_gt_zero apply (auto simp: Im_exp) using le_less apply fastforce using not_le by blast finally show ?thesis by blast qed auto lemma Arg2pi_eq_0: "Arg2pi z = 0 \ z \ \ \ 0 \ Re z" proof (cases "z=0") case False have "z \ \ \ 0 \ Re z \ z \ \ \ 0 \ Re (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... \ z \ \ \ 0 \ Re (exp (\ * complex_of_real (Arg2pi z)))" using False by (simp add: zero_le_mult_iff) also have "... \ Arg2pi z = 0" proof - have [simp]: "Arg2pi z = 0 \ z \ \" using Arg2pi_eq [of z] by (auto simp: Reals_def) moreover have "\z \ \; 0 \ cos (Arg2pi z)\ \ Arg2pi z = 0" by (metis Arg2pi_lt_pi Arg2pi_ge_0 Arg2pi_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl) ultimately show ?thesis by (auto simp: Re_exp) qed finally show ?thesis by blast qed auto corollary\<^marker>\tag unimportant\ Arg2pi_gt_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg2pi z > 0" using Arg2pi_eq_0 Arg2pi_ge_0 assms dual_order.strict_iff_order unfolding nonneg_Reals_def by fastforce lemma Arg2pi_eq_pi: "Arg2pi z = pi \ z \ \ \ Re z < 0" using Arg2pi_le_pi [of z] Arg2pi_lt_pi [of z] Arg2pi_eq_0 [of z] Arg2pi_ge_0 [of z] by (fastforce simp: complex_is_Real_iff) lemma Arg2pi_eq_0_pi: "Arg2pi z = 0 \ Arg2pi z = pi \ z \ \" using Arg2pi_eq_0 Arg2pi_eq_pi not_le by auto lemma Arg2pi_of_real: "Arg2pi (of_real r) = (if r<0 then pi else 0)" using Arg2pi_eq_0_pi Arg2pi_eq_pi by fastforce lemma Arg2pi_real: "z \ \ \ Arg2pi z = (if 0 \ Re z then 0 else pi)" using Arg2pi_eq_0 Arg2pi_eq_0_pi by auto lemma Arg2pi_inverse: "Arg2pi(inverse z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" proof (cases "z=0") case False show ?thesis apply (rule Arg2pi_unique [of "inverse (norm z)"]) using Arg2pi_eq False Arg2pi_ge_0 [of z] Arg2pi_lt_2pi [of z] Arg2pi_eq_0 [of z] by (auto simp: Arg2pi_real in_Reals_norm exp_diff field_simps) qed auto lemma Arg2pi_eq_iff: assumes "w \ 0" "z \ 0" shows "Arg2pi w = Arg2pi z \ (\x. 0 < x & w = of_real x * z)" using assms Arg2pi_eq [of z] Arg2pi_eq [of w] apply auto apply (rule_tac x="norm w / norm z" in exI) apply (simp add: field_split_simps) by (metis mult.commute mult.left_commute) lemma Arg2pi_inverse_eq_0: "Arg2pi(inverse z) = 0 \ Arg2pi z = 0" by (metis Arg2pi_eq_0 Arg2pi_inverse inverse_inverse_eq) lemma Arg2pi_divide: assumes "w \ 0" "z \ 0" "Arg2pi w \ Arg2pi z" shows "Arg2pi(z / w) = Arg2pi z - Arg2pi w" apply (rule Arg2pi_unique [of "norm(z / w)"]) using assms Arg2pi_eq Arg2pi_ge_0 [of w] Arg2pi_lt_2pi [of z] apply (auto simp: exp_diff norm_divide field_simps) done lemma Arg2pi_le_div_sum: assumes "w \ 0" "z \ 0" "Arg2pi w \ Arg2pi z" shows "Arg2pi z = Arg2pi w + Arg2pi(z / w)" by (simp add: Arg2pi_divide assms) lemma Arg2pi_le_div_sum_eq: assumes "w \ 0" "z \ 0" shows "Arg2pi w \ Arg2pi z \ Arg2pi z = Arg2pi w + Arg2pi(z / w)" using assms by (auto simp: Arg2pi_ge_0 intro: Arg2pi_le_div_sum) lemma Arg2pi_diff: assumes "w \ 0" "z \ 0" shows "Arg2pi w - Arg2pi z = (if Arg2pi z \ Arg2pi w then Arg2pi(w / z) else Arg2pi(w/z) - 2*pi)" using assms Arg2pi_divide Arg2pi_inverse [of "w/z"] Arg2pi_eq_0_pi by (force simp add: Arg2pi_ge_0 Arg2pi_divide not_le split: if_split_asm) lemma Arg2pi_add: assumes "w \ 0" "z \ 0" shows "Arg2pi w + Arg2pi z = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi(w * z) else Arg2pi(w * z) + 2*pi)" using assms Arg2pi_diff [of "w*z" z] Arg2pi_le_div_sum_eq [of z "w*z"] apply (auto simp: Arg2pi_ge_0 Arg2pi_divide not_le) apply (metis Arg2pi_lt_2pi add.commute) apply (metis (no_types) Arg2pi add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less) done lemma Arg2pi_times: assumes "w \ 0" "z \ 0" shows "Arg2pi (w * z) = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi w + Arg2pi z else (Arg2pi w + Arg2pi z) - 2*pi)" using Arg2pi_add [OF assms] by auto lemma Arg2pi_cnj_eq_inverse: "z\0 \ Arg2pi (cnj z) = Arg2pi (inverse z)" apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric] mult.commute) by (metis of_real_power zero_less_norm_iff zero_less_power) lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" proof (cases "z=0") case False then show ?thesis by (simp add: Arg2pi_cnj_eq_inverse Arg2pi_inverse) qed auto lemma Arg2pi_exp: "0 \ Im z \ Im z < 2*pi \ Arg2pi(exp z) = Im z" by (rule Arg2pi_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) lemma complex_split_polar: obtains r a::real where "z = complex_of_real r * (cos a + \ * sin a)" "0 \ r" "0 \ a" "a < 2*pi" using Arg2pi cis.ctr cis_conv_exp unfolding Complex_eq is_Arg_def by fastforce lemma Re_Im_le_cmod: "Im w * sin \ + Re w * cos \ \ cmod w" proof (cases w rule: complex_split_polar) case (1 r a) with sin_cos_le1 [of a \] show ?thesis apply (simp add: norm_mult cmod_unit_one) by (metis (no_types, hide_lams) abs_le_D1 distrib_left mult.commute mult.left_commute mult_left_le) qed subsection\<^marker>\tag unimportant\\Analytic properties of tangent function\ lemma cnj_tan: "cnj(tan z) = tan(cnj z)" by (simp add: cnj_cos cnj_sin tan_def) lemma field_differentiable_at_tan: "cos z \ 0 \ tan field_differentiable at z" unfolding field_differentiable_def using DERIV_tan by blast lemma field_differentiable_within_tan: "cos z \ 0 \ tan field_differentiable (at z within s)" using field_differentiable_at_tan field_differentiable_at_within by blast lemma continuous_within_tan: "cos z \ 0 \ continuous (at z within s) tan" using continuous_at_imp_continuous_within isCont_tan by blast lemma continuous_on_tan [continuous_intros]: "(\z. z \ s \ cos z \ 0) \ continuous_on s tan" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_tan: "(\z. z \ s \ cos z \ 0) \ tan holomorphic_on s" by (simp add: field_differentiable_within_tan holomorphic_on_def) subsection\The principal branch of the Complex logarithm\ instantiation complex :: ln begin definition\<^marker>\tag important\ ln_complex :: "complex \ complex" where "ln_complex \ \z. THE w. exp w = z & -pi < Im(w) & Im(w) \ pi" text\NOTE: within this scope, the constant Ln is not yet available!\ lemma assumes "z \ 0" shows exp_Ln [simp]: "exp(ln z) = z" and mpi_less_Im_Ln: "-pi < Im(ln z)" and Im_Ln_le_pi: "Im(ln z) \ pi" proof - obtain \ where z: "z / (cmod z) = Complex (cos \) (sin \)" using complex_unimodular_polar [of "z / (norm z)"] assms by (auto simp: norm_divide field_split_simps) obtain \ where \: "- pi < \" "\ \ pi" "sin \ = sin \" "cos \ = cos \" using sincos_principal_value [of "\"] assms by (auto simp: norm_divide field_split_simps) have "exp(ln z) = z & -pi < Im(ln z) & Im(ln z) \ pi" unfolding ln_complex_def apply (rule theI [where a = "Complex (ln(norm z)) \"]) using z assms \ apply (auto simp: field_simps exp_complex_eqI exp_eq_polar cis.code) done then show "exp(ln z) = z" "-pi < Im(ln z)" "Im(ln z) \ pi" by auto qed lemma Ln_exp [simp]: assumes "-pi < Im(z)" "Im(z) \ pi" shows "ln(exp z) = z" apply (rule exp_complex_eqI) using assms mpi_less_Im_Ln [of "exp z"] Im_Ln_le_pi [of "exp z"] apply auto done subsection\<^marker>\tag unimportant\\Relation to Real Logarithm\ lemma Ln_of_real: assumes "0 < z" shows "ln(of_real z::complex) = of_real(ln z)" proof - have "ln(of_real (exp (ln z))::complex) = ln (exp (of_real (ln z)))" by (simp add: exp_of_real) also have "... = of_real(ln z)" using assms by (subst Ln_exp) auto finally show ?thesis using assms by simp qed corollary\<^marker>\tag unimportant\ Ln_in_Reals [simp]: "z \ \ \ Re z > 0 \ ln z \ \" by (auto simp: Ln_of_real elim: Reals_cases) corollary\<^marker>\tag unimportant\ Im_Ln_of_real [simp]: "r > 0 \ Im (ln (of_real r)) = 0" by (simp add: Ln_of_real) lemma cmod_Ln_Reals [simp]: "z \ \ \ 0 < Re z \ cmod (ln z) = norm (ln (Re z))" using Ln_of_real by force lemma Ln_Reals_eq: "\x \ \; Re x > 0\ \ ln x = of_real (ln (Re x))" using Ln_of_real by force lemma Ln_1 [simp]: "ln 1 = (0::complex)" proof - have "ln (exp 0) = (0::complex)" by (simp add: del: exp_zero) then show ?thesis by simp qed lemma Ln_eq_zero_iff [simp]: "x \ \\<^sub>\\<^sub>0 \ ln x = 0 \ x = 1" for x::complex by auto (metis exp_Ln exp_zero nonpos_Reals_zero_I) instance by intro_classes (rule ln_complex_def Ln_1) end abbreviation Ln :: "complex \ complex" where "Ln \ ln" lemma Ln_eq_iff: "w \ 0 \ z \ 0 \ (Ln w = Ln z \ w = z)" by (metis exp_Ln) lemma Ln_unique: "exp(z) = w \ -pi < Im(z) \ Im(z) \ pi \ Ln w = z" using Ln_exp by blast lemma Re_Ln [simp]: "z \ 0 \ Re(Ln z) = ln(norm z)" by (metis exp_Ln ln_exp norm_exp_eq_Re) corollary\<^marker>\tag unimportant\ ln_cmod_le: assumes z: "z \ 0" shows "ln (cmod z) \ cmod (Ln z)" using norm_exp [of "Ln z", simplified exp_Ln [OF z]] by (metis Re_Ln complex_Re_le_cmod z) proposition\<^marker>\tag unimportant\ exists_complex_root: fixes z :: complex assumes "n \ 0" obtains w where "z = w ^ n" proof (cases "z=0") case False then show ?thesis by (rule_tac w = "exp(Ln z / n)" in that) (simp add: assms exp_of_nat_mult [symmetric]) qed (use assms in auto) corollary\<^marker>\tag unimportant\ exists_complex_root_nonzero: fixes z::complex assumes "z \ 0" "n \ 0" obtains w where "w \ 0" "z = w ^ n" by (metis exists_complex_root [of n z] assms power_0_left) subsection\<^marker>\tag unimportant\\Derivative of Ln away from the branch cut\ lemma assumes "z \ \\<^sub>\\<^sub>0" shows has_field_derivative_Ln: "(Ln has_field_derivative inverse(z)) (at z)" and Im_Ln_less_pi: "Im (Ln z) < pi" proof - have znz [simp]: "z \ 0" using assms by auto then have "Im (Ln z) \ pi" by (metis (no_types) Im_exp Ln_in_Reals assms complex_nonpos_Reals_iff complex_is_Real_iff exp_Ln mult_zero_right not_less pi_neq_zero sin_pi znz) then show *: "Im (Ln z) < pi" using assms Im_Ln_le_pi by (simp add: le_neq_trans) let ?U = "{w. -pi < Im(w) \ Im(w) < pi}" have 1: "open ?U" by (simp add: open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt) have 2: "\x. x \ ?U \ (exp has_derivative blinfun_apply(Blinfun ((*) (exp x)))) (at x)" apply (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right) using DERIV_exp has_field_derivative_def by blast have 3: "continuous_on ?U (\x. Blinfun ((*) (exp x)))" unfolding blinfun_mult_right.abs_eq [symmetric] by (intro continuous_intros) have 4: "Ln z \ ?U" by (auto simp: mpi_less_Im_Ln *) have 5: "Blinfun ((*) (inverse z)) o\<^sub>L Blinfun ((*) (exp (Ln z))) = id_blinfun" by (rule blinfun_eqI) (simp add: bounded_linear_mult_right bounded_linear_Blinfun_apply) obtain U' V g g' where "open U'" and sub: "U' \ ?U" and "Ln z \ U'" "open V" "z \ V" and hom: "homeomorphism U' V exp g" and g: "\y. y \ V \ (g has_derivative (g' y)) (at y)" and g': "\y. y \ V \ g' y = inv ((*) (exp (g y)))" and bij: "\y. y \ V \ bij ((*) (exp (g y)))" using inverse_function_theorem [OF 1 2 3 4 5] by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right) blast show "(Ln has_field_derivative inverse(z)) (at z)" unfolding has_field_derivative_def proof (rule has_derivative_transform_within_open) show g_eq_Ln: "g y = Ln y" if "y \ V" for y proof - obtain x where "y = exp x" "x \ U'" using hom homeomorphism_image1 that \y \ V\ by blast then show ?thesis using sub hom homeomorphism_apply1 by fastforce qed have "0 \ V" by (meson exp_not_eq_zero hom homeomorphism_def) then have "\y. y \ V \ g' y = inv ((*) y)" by (metis exp_Ln g' g_eq_Ln) then have g': "g' z = (\x. x/z)" by (metis (no_types, hide_lams) bij \z \ V\ bij_inv_eq_iff exp_Ln g_eq_Ln nonzero_mult_div_cancel_left znz) show "(g has_derivative (*) (inverse z)) (at z)" using g [OF \z \ V\] g' by (simp add: \z \ V\ field_class.field_divide_inverse has_derivative_imp_has_field_derivative has_field_derivative_imp_has_derivative) qed (auto simp: \z \ V\ \open V\) qed declare has_field_derivative_Ln [derivative_intros] declare has_field_derivative_Ln [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Ln: "z \ \\<^sub>\\<^sub>0 \ Ln field_differentiable at z" using field_differentiable_def has_field_derivative_Ln by blast lemma field_differentiable_within_Ln: "z \ \\<^sub>\\<^sub>0 \ Ln field_differentiable (at z within S)" using field_differentiable_at_Ln field_differentiable_within_subset by blast lemma continuous_at_Ln: "z \ \\<^sub>\\<^sub>0 \ continuous (at z) Ln" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Ln) lemma isCont_Ln' [simp,continuous_intros]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. Ln (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_Ln]) lemma continuous_within_Ln [continuous_intros]: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Ln" using continuous_at_Ln continuous_at_imp_continuous_within by blast lemma continuous_on_Ln [continuous_intros]: "(\z. z \ S \ z \ \\<^sub>\\<^sub>0) \ continuous_on S Ln" by (simp add: continuous_at_imp_continuous_on continuous_within_Ln) lemma continuous_on_Ln' [continuous_intros]: "continuous_on S f \ (\z. z \ S \ f z \ \\<^sub>\\<^sub>0) \ continuous_on S (\x. Ln (f x))" by (rule continuous_on_compose2[OF continuous_on_Ln, of "UNIV - nonpos_Reals" S f]) auto lemma holomorphic_on_Ln [holomorphic_intros]: "(\z. z \ S \ z \ \\<^sub>\\<^sub>0) \ Ln holomorphic_on S" by (simp add: field_differentiable_within_Ln holomorphic_on_def) lemma holomorphic_on_Ln' [holomorphic_intros]: "(\z. z \ A \ f z \ \\<^sub>\\<^sub>0) \ f holomorphic_on A \ (\z. Ln (f z)) holomorphic_on A" using holomorphic_on_compose_gen[OF _ holomorphic_on_Ln, of f A "- \\<^sub>\\<^sub>0"] by (auto simp: o_def) lemma tendsto_Ln [tendsto_intros]: fixes L F assumes "(f \ L) F" "L \ \\<^sub>\\<^sub>0" shows "((\x. Ln (f x)) \ Ln L) F" proof - have "nhds L \ filtermap f F" using assms(1) by (simp add: filterlim_def) moreover have "\\<^sub>F y in nhds L. y \ - \\<^sub>\\<^sub>0" using eventually_nhds_in_open[of "- \\<^sub>\\<^sub>0" L] assms by (auto simp: open_Compl) ultimately have "\\<^sub>F y in filtermap f F. y \ - \\<^sub>\\<^sub>0" by (rule filter_leD) moreover have "continuous_on (-\\<^sub>\\<^sub>0) Ln" by (rule continuous_on_Ln) auto ultimately show ?thesis using continuous_on_tendsto_compose[of "- \\<^sub>\\<^sub>0" Ln f L F] assms by (simp add: eventually_filtermap) qed lemma divide_ln_mono: fixes x y::real assumes "3 \ x" "x \ y" shows "x / ln x \ y / ln y" proof (rule exE [OF complex_mvt_line [of x y "\z. z / Ln z" "\z. 1/(Ln z) - 1/(Ln z)^2"]]; clarsimp simp add: closed_segment_Reals closed_segment_eq_real_ivl assms) show "\u. \x \ u; u \ y\ \ ((\z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)" using \3 \ x\ by (force intro!: derivative_eq_intros simp: field_simps power_eq_if) show "x / ln x \ y / ln y" if "Re (y / Ln y) - Re (x / Ln x) = (Re (1 / Ln u) - Re (1 / (Ln u)\<^sup>2)) * (y - x)" and x: "x \ u" "u \ y" for u proof - have eq: "y / ln y = (1 / ln u - 1 / (ln u)\<^sup>2) * (y - x) + x / ln x" using that \3 \ x\ by (auto simp: Ln_Reals_eq in_Reals_norm group_add_class.diff_eq_eq) show ?thesis using exp_le \3 \ x\ x by (simp add: eq) (simp add: power_eq_if divide_simps ln_ge_iff) qed qed theorem Ln_series: fixes z :: complex assumes "norm z < 1" shows "(\n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\n. ?f n * z^n) sums _") proof - let ?F = "\z. \n. ?f n * z^n" and ?F' = "\z. \n. diffs ?f n * z^n" have r: "conv_radius ?f = 1" by (intro conv_radius_ratio_limit_nonzero[of _ 1]) (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc) have "\c. \z\ball 0 1. ln (1 + z) - ?F z = c" proof (rule has_field_derivative_zero_constant) fix z :: complex assume z': "z \ ball 0 1" hence z: "norm z < 1" by simp define t :: complex where "t = of_real (1 + norm z) / 2" from z have t: "norm z < norm t" "norm t < 1" unfolding t_def by (simp_all add: field_simps norm_divide del: of_real_add) have "Re (-z) \ norm (-z)" by (rule complex_Re_le_cmod) also from z have "... < 1" by simp finally have "((\z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)" by (auto intro!: derivative_eq_intros simp: complex_nonpos_Reals_iff) moreover have "(?F has_field_derivative ?F' z) (at z)" using t r by (intro termdiffs_strong[of _ t] summable_in_conv_radius) simp_all ultimately have "((\z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z)) (at z within ball 0 1)" by (intro derivative_intros) (simp_all add: at_within_open[OF z']) also have "(\n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all from sums_split_initial_segment[OF this, of 1] have "(\i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc) hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse) also have "inverse (1 + z) - inverse (1 + z) = 0" by simp finally show "((\z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" . qed simp_all then obtain c where c: "\z. z \ ball 0 1 \ ln (1 + z) - ?F z = c" by blast from c[of 0] have "c = 0" by (simp only: powser_zero) simp with c[of z] assms have "ln (1 + z) = ?F z" by simp moreover have "summable (\n. ?f n * z^n)" using assms r by (intro summable_in_conv_radius) simp_all ultimately show ?thesis by (simp add: sums_iff) qed lemma Ln_series': "cmod z < 1 \ (\n. - ((-z)^n) / of_nat n) sums ln (1 + z)" by (drule Ln_series) (simp add: power_minus') lemma ln_series': assumes "abs (x::real) < 1" shows "(\n. - ((-x)^n) / of_nat n) sums ln (1 + x)" proof - from assms have "(\n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)" by (intro Ln_series') simp_all also have "(\n. - ((-of_real x)^n) / of_nat n) = (\n. complex_of_real (- ((-x)^n) / of_nat n))" by (rule ext) simp also from assms have "ln (1 + complex_of_real x) = of_real (ln (1 + x))" by (subst Ln_of_real [symmetric]) simp_all finally show ?thesis by (subst (asm) sums_of_real_iff) qed lemma Ln_approx_linear: fixes z :: complex assumes "norm z < 1" shows "norm (ln (1 + z) - z) \ norm z^2 / (1 - norm z)" proof - let ?f = "\n. (-1)^Suc n / of_nat n" from assms have "(\n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp moreover have "(\n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp ultimately have "(\n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)" by (subst left_diff_distrib, intro sums_diff) simp_all from sums_split_initial_segment[OF this, of "Suc 1"] have "(\i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)" by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse) hence "(Ln (1 + z) - z) = (\i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)" by (simp add: sums_iff) also have A: "summable (\n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))" by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]]) (auto simp: assms field_simps intro!: always_eventually) hence "norm (\i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i) \ (\i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))" by (intro summable_norm) (auto simp: norm_power norm_inverse norm_mult mult_ac simp del: of_nat_add of_nat_Suc) also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \ norm ((-z)^2 * (-z)^i) * 1" for i by (intro mult_left_mono) (simp_all add: field_split_simps) hence "(\i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i)) \ (\i. norm (-(z^2) * (-z)^i))" using A assms apply (simp_all only: norm_power norm_inverse norm_divide norm_mult) apply (intro suminf_le summable_mult summable_geometric) apply (auto simp: norm_power field_simps simp del: of_nat_add of_nat_Suc) done also have "... = norm z^2 * (\i. norm z^i)" using assms by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power) also have "(\i. norm z^i) = inverse (1 - norm z)" using assms by (subst suminf_geometric) (simp_all add: divide_inverse) also have "norm z^2 * ... = norm z^2 / (1 - norm z)" by (simp add: divide_inverse) finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Quadrant-type results for Ln\ lemma cos_lt_zero_pi: "pi/2 < x \ x < 3*pi/2 \ cos x < 0" using cos_minus_pi cos_gt_zero_pi [of "x-pi"] by simp lemma Re_Ln_pos_lt: assumes "z \ 0" shows "\Im(Ln z)\ < pi/2 \ 0 < Re(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "\Im w\ < pi/2 \ 0 < Re(exp w)" using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"] apply (auto simp: Re_exp zero_less_mult_iff cos_gt_zero_pi abs_if split: if_split_asm) apply (metis cos_minus cos_pi_half divide_minus_left less_irrefl linorder_neqE_linordered_idom nonzero_mult_div_cancel_right zero_neq_numeral)+ done } then show ?thesis using assms by auto qed lemma Re_Ln_pos_le: assumes "z \ 0" shows "\Im(Ln z)\ \ pi/2 \ 0 \ Re(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "\Im w\ \ pi/2 \ 0 \ Re(exp w)" apply (auto simp: Re_exp zero_le_mult_iff cos_ge_zero) using cos_lt_zero_pi [of "- (Im w)"] cos_lt_zero_pi [of "(Im w)"] not_le apply (auto simp: abs_if split: if_split_asm) done } then show ?thesis using assms by auto qed lemma Im_Ln_pos_lt: assumes "z \ 0" shows "0 < Im(Ln z) \ Im(Ln z) < pi \ 0 < Im(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "0 < Im w \ Im w < pi \ 0 < Im(exp w)" using sin_gt_zero [of "- (Im w)"] sin_gt_zero [of "(Im w)"] apply (simp add: Im_exp zero_less_mult_iff) using less_linear apply fastforce done } then show ?thesis using assms by auto qed lemma Im_Ln_pos_le: assumes "z \ 0" shows "0 \ Im(Ln z) \ Im(Ln z) \ pi \ 0 \ Im(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "0 \ Im w \ Im w \ pi \ 0 \ Im(exp w)" using sin_ge_zero [of "- (Im w)"] sin_ge_zero [of "(Im w)"] apply (auto simp: Im_exp zero_le_mult_iff sin_ge_zero) apply (metis not_le not_less_iff_gr_or_eq pi_not_less_zero sin_eq_0_pi) done } then show ?thesis using assms by auto qed lemma Re_Ln_pos_lt_imp: "0 < Re(z) \ \Im(Ln z)\ < pi/2" by (metis Re_Ln_pos_lt less_irrefl zero_complex.simps(1)) lemma Im_Ln_pos_lt_imp: "0 < Im(z) \ 0 < Im(Ln z) \ Im(Ln z) < pi" by (metis Im_Ln_pos_lt not_le order_refl zero_complex.simps(2)) text\A reference to the set of positive real numbers\ lemma Im_Ln_eq_0: "z \ 0 \ (Im(Ln z) = 0 \ 0 < Re(z) \ Im(z) = 0)" by (metis Im_complex_of_real Im_exp Ln_in_Reals Re_Ln_pos_lt Re_Ln_pos_lt_imp Re_complex_of_real complex_is_Real_iff exp_Ln exp_of_real pi_gt_zero) lemma Im_Ln_eq_pi: "z \ 0 \ (Im(Ln z) = pi \ Re(z) < 0 \ Im(z) = 0)" by (metis Im_Ln_eq_0 Im_Ln_pos_le Im_Ln_pos_lt add.left_neutral complex_eq less_eq_real_def mult_zero_right not_less_iff_gr_or_eq pi_ge_zero pi_neq_zero rcis_zero_arg rcis_zero_mod) subsection\<^marker>\tag unimportant\\More Properties of Ln\ lemma cnj_Ln: assumes "z \ \\<^sub>\\<^sub>0" shows "cnj(Ln z) = Ln(cnj z)" proof (cases "z=0") case False show ?thesis proof (rule exp_complex_eqI) have "\Im (cnj (Ln z)) - Im (Ln (cnj z))\ \ \Im (cnj (Ln z))\ + \Im (Ln (cnj z))\" by (rule abs_triangle_ineq4) also have "... < pi + pi" proof - have "\Im (cnj (Ln z))\ < pi" by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln) moreover have "\Im (Ln (cnj z))\ \ pi" by (meson abs_le_iff complex_cnj_zero_iff less_eq_real_def minus_less_iff False Im_Ln_le_pi mpi_less_Im_Ln) ultimately show ?thesis by simp qed finally show "\Im (cnj (Ln z)) - Im (Ln (cnj z))\ < 2 * pi" by simp show "exp (cnj (Ln z)) = exp (Ln (cnj z))" by (metis False complex_cnj_zero_iff exp_Ln exp_cnj) qed qed (use assms in auto) lemma Ln_inverse: assumes "z \ \\<^sub>\\<^sub>0" shows "Ln(inverse z) = -(Ln z)" proof (cases "z=0") case False show ?thesis proof (rule exp_complex_eqI) have "\Im (Ln (inverse z)) - Im (- Ln z)\ \ \Im (Ln (inverse z))\ + \Im (- Ln z)\" by (rule abs_triangle_ineq4) also have "... < pi + pi" proof - have "\Im (Ln (inverse z))\ < pi" by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln) moreover have "\Im (- Ln z)\ \ pi" using False Im_Ln_le_pi mpi_less_Im_Ln by fastforce ultimately show ?thesis by simp qed finally show "\Im (Ln (inverse z)) - Im (- Ln z)\ < 2 * pi" by simp show "exp (Ln (inverse z)) = exp (- Ln z)" by (simp add: False exp_minus) qed qed (use assms in auto) lemma Ln_minus1 [simp]: "Ln(-1) = \ * pi" apply (rule exp_complex_eqI) using Im_Ln_le_pi [of "-1"] mpi_less_Im_Ln [of "-1"] cis_conv_exp cis_pi apply (auto simp: abs_if) done lemma Ln_ii [simp]: "Ln \ = \ * of_real pi/2" using Ln_exp [of "\ * (of_real pi/2)"] unfolding exp_Euler by simp lemma Ln_minus_ii [simp]: "Ln(-\) = - (\ * pi/2)" proof - have "Ln(-\) = Ln(inverse \)" by simp also have "... = - (Ln \)" using Ln_inverse by blast also have "... = - (\ * pi/2)" by simp finally show ?thesis . qed lemma Ln_times: assumes "w \ 0" "z \ 0" shows "Ln(w * z) = (if Im(Ln w + Ln z) \ -pi then (Ln(w) + Ln(z)) + \ * of_real(2*pi) else if Im(Ln w + Ln z) > pi then (Ln(w) + Ln(z)) - \ * of_real(2*pi) else Ln(w) + Ln(z))" using pi_ge_zero Im_Ln_le_pi [of w] Im_Ln_le_pi [of z] using assms mpi_less_Im_Ln [of w] mpi_less_Im_Ln [of z] by (auto simp: exp_add exp_diff sin_double cos_double exp_Euler intro!: Ln_unique) corollary\<^marker>\tag unimportant\ Ln_times_simple: "\w \ 0; z \ 0; -pi < Im(Ln w) + Im(Ln z); Im(Ln w) + Im(Ln z) \ pi\ \ Ln(w * z) = Ln(w) + Ln(z)" by (simp add: Ln_times) corollary\<^marker>\tag unimportant\ Ln_times_of_real: "\r > 0; z \ 0\ \ Ln(of_real r * z) = ln r + Ln(z)" using mpi_less_Im_Ln Im_Ln_le_pi by (force simp: Ln_times) corollary\<^marker>\tag unimportant\ Ln_times_Reals: "\r \ Reals; Re r > 0; z \ 0\ \ Ln(r * z) = ln (Re r) + Ln(z)" using Ln_Reals_eq Ln_times_of_real by fastforce corollary\<^marker>\tag unimportant\ Ln_divide_of_real: "\r > 0; z \ 0\ \ Ln(z / of_real r) = Ln(z) - ln r" using Ln_times_of_real [of "inverse r" z] by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric] del: of_real_inverse) corollary\<^marker>\tag unimportant\ Ln_prod: fixes f :: "'a \ complex" assumes "finite A" "\x. x \ A \ f x \ 0" shows "\n. Ln (prod f A) = (\x \ A. Ln (f x) + (of_int (n x) * (2 * pi)) * \)" using assms proof (induction A) case (insert x A) then obtain n where n: "Ln (prod f A) = (\x\A. Ln (f x) + of_real (of_int (n x) * (2 * pi)) * \)" by auto define D where "D \ Im (Ln (f x)) + Im (Ln (prod f A))" define q::int where "q \ (if D \ -pi then 1 else if D > pi then -1 else 0)" have "prod f A \ 0" "f x \ 0" by (auto simp: insert.hyps insert.prems) with insert.hyps pi_ge_zero show ?case by (rule_tac x="n(x:=q)" in exI) (force simp: Ln_times q_def D_def n intro!: sum.cong) qed auto lemma Ln_minus: assumes "z \ 0" shows "Ln(-z) = (if Im(z) \ 0 \ \(Re(z) < 0 \ Im(z) = 0) then Ln(z) + \ * pi else Ln(z) - \ * pi)" (is "_ = ?rhs") using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] by (fastforce simp: exp_add exp_diff exp_Euler intro!: Ln_unique) lemma Ln_inverse_if: assumes "z \ 0" shows "Ln (inverse z) = (if z \ \\<^sub>\\<^sub>0 then -(Ln z) + \ * 2 * complex_of_real pi else -(Ln z))" proof (cases "z \ \\<^sub>\\<^sub>0") case False then show ?thesis by (simp add: Ln_inverse) next case True then have z: "Im z = 0" "Re z < 0" using assms apply (auto simp: complex_nonpos_Reals_iff) by (metis complex_is_Real_iff le_imp_less_or_eq of_real_0 of_real_Re) have "Ln(inverse z) = Ln(- (inverse (-z)))" by simp also have "... = Ln (inverse (-z)) + \ * complex_of_real pi" using assms z apply (simp add: Ln_minus) apply (simp add: field_simps) done also have "... = - Ln (- z) + \ * complex_of_real pi" apply (subst Ln_inverse) using z by (auto simp add: complex_nonneg_Reals_iff) also have "... = - (Ln z) + \ * 2 * complex_of_real pi" by (subst Ln_minus) (use assms z in auto) finally show ?thesis by (simp add: True) qed lemma Ln_times_ii: assumes "z \ 0" shows "Ln(\ * z) = (if 0 \ Re(z) | Im(z) < 0 then Ln(z) + \ * of_real pi/2 else Ln(z) - \ * of_real(3 * pi/2))" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] Re_Ln_pos_le [of z] by (simp add: Ln_times) auto lemma Ln_of_nat [simp]: "0 < n \ Ln (of_nat n) = of_real (ln (of_nat n))" by (subst of_real_of_nat_eq[symmetric], subst Ln_of_real[symmetric]) simp_all lemma Ln_of_nat_over_of_nat: assumes "m > 0" "n > 0" shows "Ln (of_nat m / of_nat n) = of_real (ln (of_nat m) - ln (of_nat n))" proof - have "of_nat m / of_nat n = (of_real (of_nat m / of_nat n) :: complex)" by simp also from assms have "Ln ... = of_real (ln (of_nat m / of_nat n))" by (simp add: Ln_of_real[symmetric]) also from assms have "... = of_real (ln (of_nat m) - ln (of_nat n))" by (simp add: ln_div) finally show ?thesis . qed subsection\The Argument of a Complex Number\ text\Finally: it's is defined for the same interval as the complex logarithm: \(-\,\]\.\ definition\<^marker>\tag important\ Arg :: "complex \ real" where "Arg z \ (if z = 0 then 0 else Im (Ln z))" lemma Arg_of_real: "Arg (of_real r) = (if r<0 then pi else 0)" by (simp add: Im_Ln_eq_pi Arg_def) lemma mpi_less_Arg: "-pi < Arg z" and Arg_le_pi: "Arg z \ pi" by (auto simp: Arg_def mpi_less_Im_Ln Im_Ln_le_pi) lemma assumes "z \ 0" shows Arg_eq: "z = of_real(norm z) * exp(\ * Arg z)" using assms exp_Ln exp_eq_polar by (auto simp: Arg_def) lemma is_Arg_Arg: "z \ 0 \ is_Arg z (Arg z)" by (simp add: Arg_eq is_Arg_def) lemma Argument_exists: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" obtains s where "is_Arg z s" "s\R" proof - let ?rp = "r - Arg z + pi" define k where "k \ \?rp / (2 * pi)\" have "(Arg z + of_int k * (2 * pi)) \ R" using floor_divide_lower [of "2*pi" ?rp] floor_divide_upper [of "2*pi" ?rp] by (auto simp: k_def algebra_simps R) then show ?thesis using Arg_eq \z \ 0\ is_Arg_2pi_iff is_Arg_def that by blast qed lemma Argument_exists_unique: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" obtains s where "is_Arg z s" "s\R" "\t. \is_Arg z t; t\R\ \ s=t" proof - obtain s where s: "is_Arg z s" "s\R" using Argument_exists [OF assms] . moreover have "\t. \is_Arg z t; t\R\ \ s=t" using assms s by (auto simp: is_Arg_eqI) ultimately show thesis using that by blast qed lemma Argument_Ex1: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" shows "\!s. is_Arg z s \ s \ R" using Argument_exists_unique [OF assms] by metis lemma Arg_divide: assumes "w \ 0" "z \ 0" shows "is_Arg (z / w) (Arg z - Arg w)" using Arg_eq [of z] Arg_eq [of w] Arg_eq [of "norm(z / w)"] assms by (auto simp: is_Arg_def norm_divide field_simps exp_diff Arg_of_real) lemma Arg_unique_lemma: assumes z: "is_Arg z t" and z': "is_Arg z t'" and t: "- pi < t" "t \ pi" and t': "- pi < t'" "t' \ pi" and nz: "z \ 0" shows "t' = t" using Arg2pi_unique_lemma [of "- (inverse z)"] proof - have "pi - t' = pi - t" proof (rule Arg2pi_unique_lemma [of "- (inverse z)"]) have "- (inverse z) = - (inverse (of_real(norm z) * exp(\ * t)))" by (metis is_Arg_def z) also have "... = (cmod (- inverse z)) * exp (\ * (pi - t))" by (auto simp: field_simps exp_diff norm_divide) finally show "is_Arg (- inverse z) (pi - t)" unfolding is_Arg_def . have "- (inverse z) = - (inverse (of_real(norm z) * exp(\ * t')))" by (metis is_Arg_def z') also have "... = (cmod (- inverse z)) * exp (\ * (pi - t'))" by (auto simp: field_simps exp_diff norm_divide) finally show "is_Arg (- inverse z) (pi - t')" unfolding is_Arg_def . qed (use assms in auto) then show ?thesis by simp qed lemma complex_norm_eq_1_exp_eq: "norm z = 1 \ exp(\ * (Arg z)) = z" by (metis Arg_eq exp_not_eq_zero exp_zero mult.left_neutral norm_zero of_real_1 norm_exp_i_times) lemma Arg_unique: "\of_real r * exp(\ * a) = z; 0 < r; -pi < a; a \ pi\ \ Arg z = a" by (rule Arg_unique_lemma [unfolded is_Arg_def, OF _ Arg_eq]) (use mpi_less_Arg Arg_le_pi in \auto simp: norm_mult\) lemma Arg_minus: assumes "z \ 0" shows "Arg (-z) = (if Arg z \ 0 then Arg z + pi else Arg z - pi)" proof - have [simp]: "cmod z * cos (Arg z) = Re z" using assms Arg_eq [of z] by (metis Re_exp exp_Ln norm_exp_eq_Re Arg_def) have [simp]: "cmod z * sin (Arg z) = Im z" using assms Arg_eq [of z] by (metis Im_exp exp_Ln norm_exp_eq_Re Arg_def) show ?thesis apply (rule Arg_unique [of "norm z"]) apply (rule complex_eqI) using mpi_less_Arg [of z] Arg_le_pi [of z] assms apply (auto simp: Re_exp Im_exp cos_diff sin_diff cis_conv_exp [symmetric]) done qed lemma Arg_times_of_real [simp]: assumes "0 < r" shows "Arg (of_real r * z) = Arg z" proof (cases "z=0") case True then show ?thesis by (simp add: Arg_def) next case False with Arg_eq assms show ?thesis by (auto simp: mpi_less_Arg Arg_le_pi intro!: Arg_unique [of "r * norm z"]) qed lemma Arg_times_of_real2 [simp]: "0 < r \ Arg (z * of_real r) = Arg z" by (metis Arg_times_of_real mult.commute) lemma Arg_divide_of_real [simp]: "0 < r \ Arg (z / of_real r) = Arg z" by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) lemma Arg_less_0: "0 \ Arg z \ 0 \ Im z" using Im_Ln_le_pi Im_Ln_pos_le by (simp add: Arg_def) lemma Arg_eq_pi: "Arg z = pi \ Re z < 0 \ Im z = 0" by (auto simp: Arg_def Im_Ln_eq_pi) lemma Arg_lt_pi: "0 < Arg z \ Arg z < pi \ 0 < Im z" using Arg_less_0 [of z] Im_Ln_pos_lt by (auto simp: order.order_iff_strict Arg_def) lemma Arg_eq_0: "Arg z = 0 \ z \ \ \ 0 \ Re z" using complex_is_Real_iff by (simp add: Arg_def Im_Ln_eq_0) (metis less_eq_real_def of_real_Re of_real_def scale_zero_left) corollary\<^marker>\tag unimportant\ Arg_ne_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg z \ 0" using assms by (auto simp: nonneg_Reals_def Arg_eq_0) lemma Arg_eq_pi_iff: "Arg z = pi \ z \ \ \ Re z < 0" proof (cases "z=0") case False then show ?thesis using Arg_eq_0 [of "-z"] Arg_eq_pi complex_is_Real_iff by blast qed (simp add: Arg_def) lemma Arg_eq_0_pi: "Arg z = 0 \ Arg z = pi \ z \ \" using Arg_eq_pi_iff Arg_eq_0 by force lemma Arg_real: "z \ \ \ Arg z = (if 0 \ Re z then 0 else pi)" using Arg_eq_0 Arg_eq_0_pi by auto lemma Arg_inverse: "Arg(inverse z) = (if z \ \ then Arg z else - Arg z)" proof (cases "z \ \") case True then show ?thesis by simp (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse) next case False then have "Arg z < pi" "z \ 0" using Arg_eq_0_pi Arg_le_pi by (auto simp: less_eq_real_def) then show ?thesis apply (simp add: False) apply (rule Arg_unique [of "inverse (norm z)"]) using False mpi_less_Arg [of z] Arg_eq [of z] apply (auto simp: exp_minus field_simps) done qed lemma Arg_eq_iff: assumes "w \ 0" "z \ 0" shows "Arg w = Arg z \ (\x. 0 < x \ w = of_real x * z)" using assms Arg_eq [of z] Arg_eq [of w] apply auto apply (rule_tac x="norm w / norm z" in exI) apply (simp add: field_split_simps) by (metis mult.commute mult.left_commute) lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \ Arg z = 0" by (metis Arg_eq_0 Arg_inverse inverse_inverse_eq) lemma Arg_cnj_eq_inverse: "z\0 \ Arg (cnj z) = Arg (inverse z)" apply (simp add: Arg_eq_iff field_split_simps complex_norm_square [symmetric] mult.commute) by (metis of_real_power zero_less_norm_iff zero_less_power) lemma Arg_cnj: "Arg(cnj z) = (if z \ \ then Arg z else - Arg z)" by (metis Arg_cnj_eq_inverse Arg_inverse Reals_0 complex_cnj_zero) lemma Arg_exp: "-pi < Im z \ Im z \ pi \ Arg(exp z) = Im z" by (rule Arg_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) lemma Ln_Arg: "z\0 \ Ln(z) = ln(norm z) + \ * Arg(z)" by (metis Arg_def Re_Ln complex_eq) lemma continuous_at_Arg: assumes "z \ \\<^sub>\\<^sub>0" shows "continuous (at z) Arg" proof - have [simp]: "(\z. Im (Ln z)) \z\ Arg z" using Arg_def assms continuous_at by fastforce show ?thesis unfolding continuous_at proof (rule Lim_transform_within_open) show "\w. \w \ - \\<^sub>\\<^sub>0; w \ z\ \ Im (Ln w) = Arg w" by (metis Arg_def Compl_iff nonpos_Reals_zero_I) qed (use assms in auto) qed lemma continuous_within_Arg: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Arg" using continuous_at_Arg continuous_at_imp_continuous_within by blast subsection\The Unwinding Number and the Ln product Formula\ text\Note that in this special case the unwinding number is -1, 0 or 1. But it's always an integer.\ lemma is_Arg_exp_Im: "is_Arg (exp z) (Im z)" using exp_eq_polar is_Arg_def norm_exp_eq_Re by auto lemma is_Arg_exp_diff_2pi: assumes "is_Arg (exp z) \" shows "\k. Im z - of_int k * (2 * pi) = \" proof (intro exI is_Arg_eqI) let ?k = "\(Im z - \) / (2 * pi)\" show "is_Arg (exp z) (Im z - real_of_int ?k * (2 * pi))" by (metis diff_add_cancel is_Arg_2pi_iff is_Arg_exp_Im) show "\Im z - real_of_int ?k * (2 * pi) - \\ < 2 * pi" using floor_divide_upper [of "2*pi" "Im z - \"] floor_divide_lower [of "2*pi" "Im z - \"] by (auto simp: algebra_simps abs_if) qed (auto simp: is_Arg_exp_Im assms) lemma Arg_exp_diff_2pi: "\k. Im z - of_int k * (2 * pi) = Arg (exp z)" using is_Arg_exp_diff_2pi [OF is_Arg_Arg] by auto lemma unwinding_in_Ints: "(z - Ln(exp z)) / (of_real(2*pi) * \) \ \" using Arg_exp_diff_2pi [of z] by (force simp: Ints_def image_def field_simps Arg_def intro!: complex_eqI) definition\<^marker>\tag important\ unwinding :: "complex \ int" where "unwinding z \ THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \)" lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \)" using unwinding_in_Ints [of z] unfolding unwinding_def Ints_def by force lemma unwinding_2pi: "(2*pi) * \ * unwinding(z) = z - Ln(exp z)" by (simp add: unwinding) lemma Ln_times_unwinding: "w \ 0 \ z \ 0 \ Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * \ * unwinding(Ln w + Ln z)" using unwinding_2pi by (simp add: exp_add) subsection\<^marker>\tag unimportant\\Relation between Ln and Arg2pi, and hence continuity of Arg2pi\ lemma Arg2pi_Ln: assumes "0 < Arg2pi z" shows "Arg2pi z = Im(Ln(-z)) + pi" proof (cases "z = 0") case True with assms show ?thesis by simp next case False then have "z / of_real(norm z) = exp(\ * of_real(Arg2pi z))" using Arg2pi [of z] by (metis is_Arg_def abs_norm_cancel nonzero_mult_div_cancel_left norm_of_real zero_less_norm_iff) then have "- z / of_real(norm z) = exp (\ * (of_real (Arg2pi z) - pi))" using cis_conv_exp cis_pi by (auto simp: exp_diff algebra_simps) then have "ln (- z / of_real(norm z)) = ln (exp (\ * (of_real (Arg2pi z) - pi)))" by simp also have "... = \ * (of_real(Arg2pi z) - pi)" using Arg2pi [of z] assms pi_not_less_zero by auto finally have "Arg2pi z = Im (Ln (- z / of_real (cmod z))) + pi" by simp also have "... = Im (Ln (-z) - ln (cmod z)) + pi" by (metis diff_0_right minus_diff_eq zero_less_norm_iff Ln_divide_of_real False) also have "... = Im (Ln (-z)) + pi" by simp finally show ?thesis . qed lemma continuous_at_Arg2pi: assumes "z \ \\<^sub>\\<^sub>0" shows "continuous (at z) Arg2pi" proof - have *: "isCont (\z. Im (Ln (- z)) + pi) z" by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+ have [simp]: "Im x \ 0 \ Im (Ln (- x)) + pi = Arg2pi x" for x using Arg2pi_Ln by (simp add: Arg2pi_gt_0 complex_nonneg_Reals_iff) consider "Re z < 0" | "Im z \ 0" using assms using complex_nonneg_Reals_iff not_le by blast then have [simp]: "(\z. Im (Ln (- z)) + pi) \z\ Arg2pi z" using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within) show ?thesis unfolding continuous_at proof (rule Lim_transform_within_open) show "\x. \x \ - \\<^sub>\\<^sub>0; x \ z\ \ Im (Ln (- x)) + pi = Arg2pi x" by (auto simp add: Arg2pi_Ln [OF Arg2pi_gt_0] complex_nonneg_Reals_iff) qed (use assms in auto) qed text\Relation between Arg2pi and arctangent in upper halfplane\ lemma Arg2pi_arctan_upperhalf: assumes "0 < Im z" shows "Arg2pi z = pi/2 - arctan(Re z / Im z)" proof (cases "z = 0") case False show ?thesis proof (rule Arg2pi_unique [of "norm z"]) show "(cmod z) * exp (\ * (pi / 2 - arctan (Re z / Im z))) = z" apply (auto simp: exp_Euler cos_diff sin_diff) using assms norm_complex_def [of z, symmetric] apply (simp add: sin_of_real cos_of_real sin_arctan cos_arctan field_simps real_sqrt_divide) apply (metis complex_eq) done qed (use False arctan [of "Re z / Im z"] in auto) qed (use assms in auto) lemma Arg2pi_eq_Im_Ln: assumes "0 \ Im z" "0 < Re z" shows "Arg2pi z = Im (Ln z)" proof (cases "Im z = 0") case True then show ?thesis using Arg2pi_eq_0 Ln_in_Reals assms(2) complex_is_Real_iff by auto next case False then have *: "Arg2pi z > 0" using Arg2pi_gt_0 complex_is_Real_iff by blast then have "z \ 0" by auto with * assms False show ?thesis by (subst Arg2pi_Ln) (auto simp: Ln_minus) qed lemma continuous_within_upperhalf_Arg2pi: assumes "z \ 0" shows "continuous (at z within {z. 0 \ Im z}) Arg2pi" proof (cases "z \ \\<^sub>\\<^sub>0") case False then show ?thesis using continuous_at_Arg2pi continuous_at_imp_continuous_within by auto next case True then have z: "z \ \" "0 < Re z" using assms by (auto simp: complex_nonneg_Reals_iff complex_is_Real_iff complex_neq_0) then have [simp]: "Arg2pi z = 0" "Im (Ln z) = 0" by (auto simp: Arg2pi_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff) show ?thesis proof (clarsimp simp add: continuous_within Lim_within dist_norm) fix e::real assume "0 < e" moreover have "continuous (at z) (\x. Im (Ln x))" using z by (simp add: continuous_at_Ln complex_nonpos_Reals_iff) ultimately obtain d where d: "d>0" "\x. x \ z \ cmod (x - z) < d \ \Im (Ln x)\ < e" by (auto simp: continuous_within Lim_within dist_norm) { fix x assume "cmod (x - z) < Re z / 2" then have "\Re x - Re z\ < Re z / 2" by (metis le_less_trans abs_Re_le_cmod minus_complex.simps(1)) then have "0 < Re x" using z by linarith } then show "\d>0. \x. 0 \ Im x \ x \ z \ cmod (x - z) < d \ \Arg2pi x\ < e" apply (rule_tac x="min d (Re z / 2)" in exI) using z d apply (auto simp: Arg2pi_eq_Im_Ln) done qed qed lemma continuous_on_upperhalf_Arg2pi: "continuous_on ({z. 0 \ Im z} - {0}) Arg2pi" unfolding continuous_on_eq_continuous_within by (metis DiffE Diff_subset continuous_within_subset continuous_within_upperhalf_Arg2pi insertCI) lemma open_Arg2pi2pi_less_Int: assumes "0 \ s" "t \ 2*pi" shows "open ({y. s < Arg2pi y} \ {y. Arg2pi y < t})" proof - have 1: "continuous_on (UNIV - \\<^sub>\\<^sub>0) Arg2pi" using continuous_at_Arg2pi continuous_at_imp_continuous_within by (auto simp: continuous_on_eq_continuous_within) have 2: "open (UNIV - \\<^sub>\\<^sub>0 :: complex set)" by (simp add: open_Diff) have "open ({z. s < z} \ {z. z < t})" using open_lessThan [of t] open_greaterThan [of s] by (metis greaterThan_def lessThan_def open_Int) moreover have "{y. s < Arg2pi y} \ {y. Arg2pi y < t} \ - \\<^sub>\\<^sub>0" using assms by (auto simp: Arg2pi_real complex_nonneg_Reals_iff complex_is_Real_iff) ultimately show ?thesis using continuous_imp_open_vimage [OF 1 2, of "{z. Re z > s} \ {z. Re z < t}"] by auto qed lemma open_Arg2pi2pi_gt: "open {z. t < Arg2pi z}" proof (cases "t < 0") case True then have "{z. t < Arg2pi z} = UNIV" using Arg2pi_ge_0 less_le_trans by auto then show ?thesis by simp next case False then show ?thesis using open_Arg2pi2pi_less_Int [of t "2*pi"] Arg2pi_lt_2pi by auto qed lemma closed_Arg2pi2pi_le: "closed {z. Arg2pi z \ t}" using open_Arg2pi2pi_gt [of t] by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le) subsection\<^marker>\tag unimportant\\Complex Powers\ lemma powr_to_1 [simp]: "z powr 1 = (z::complex)" by (simp add: powr_def) lemma powr_nat: fixes n::nat and z::complex shows "z powr n = (if z = 0 then 0 else z^n)" by (simp add: exp_of_nat_mult powr_def) lemma norm_powr_real: "w \ \ \ 0 < Re w \ norm(w powr z) = exp(Re z * ln(Re w))" apply (simp add: powr_def) using Im_Ln_eq_0 complex_is_Real_iff norm_complex_def by auto lemma powr_complexpow [simp]: fixes x::complex shows "x \ 0 \ x powr (of_nat n) = x^n" by (induct n) (auto simp: ac_simps powr_add) lemma powr_complexnumeral [simp]: fixes x::complex shows "x \ 0 \ x powr (numeral n) = x ^ (numeral n)" by (metis of_nat_numeral powr_complexpow) lemma cnj_powr: assumes "Im a = 0 \ Re a \ 0" shows "cnj (a powr b) = cnj a powr cnj b" proof (cases "a = 0") case False with assms have "a \ \\<^sub>\\<^sub>0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) with False show ?thesis by (simp add: powr_def exp_cnj cnj_Ln) qed simp lemma powr_real_real: assumes "w \ \" "z \ \" "0 < Re w" shows "w powr z = exp(Re z * ln(Re w))" proof - have "w \ 0" using assms by auto with assms show ?thesis by (simp add: powr_def Ln_Reals_eq of_real_exp) qed lemma powr_of_real: fixes x::real and y::real shows "0 \ x \ of_real x powr (of_real y::complex) = of_real (x powr y)" by (simp_all add: powr_def exp_eq_polar) lemma powr_of_int: fixes z::complex and n::int assumes "z\(0::complex)" shows "z powr of_int n = (if n\0 then z^nat n else inverse (z^nat (-n)))" by (metis assms not_le of_int_of_nat powr_complexpow powr_minus) lemma powr_Reals_eq: "\x \ \; y \ \; Re x \ 0\ \ x powr y = of_real (Re x powr Re y)" by (metis of_real_Re powr_of_real) lemma norm_powr_real_mono: "\w \ \; 1 < Re w\ \ cmod(w powr z1) \ cmod(w powr z2) \ Re z1 \ Re z2" by (auto simp: powr_def algebra_simps Reals_def Ln_of_real) lemma powr_times_real: "\x \ \; y \ \; 0 \ Re x; 0 \ Re y\ \ (x * y) powr z = x powr z * y powr z" by (auto simp: Reals_def powr_def Ln_times exp_add algebra_simps less_eq_real_def Ln_of_real) lemma Re_powr_le: "r \ \\<^sub>\\<^sub>0 \ Re (r powr z) \ Re r powr Re z" by (auto simp: powr_def nonneg_Reals_def order_trans [OF complex_Re_le_cmod]) lemma fixes w::complex shows Reals_powr [simp]: "\w \ \\<^sub>\\<^sub>0; z \ \\ \ w powr z \ \" and nonneg_Reals_powr [simp]: "\w \ \\<^sub>\\<^sub>0; z \ \\ \ w powr z \ \\<^sub>\\<^sub>0" by (auto simp: nonneg_Reals_def Reals_def powr_of_real) lemma powr_neg_real_complex: shows "(- of_real x) powr a = (-1) powr (of_real (sgn x) * a) * of_real x powr (a :: complex)" proof (cases "x = 0") assume x: "x \ 0" hence "(-x) powr a = exp (a * ln (-of_real x))" by (simp add: powr_def) also from x have "ln (-of_real x) = Ln (of_real x) + of_real (sgn x) * pi * \" by (simp add: Ln_minus Ln_of_real) also from x have "exp (a * ...) = cis pi powr (of_real (sgn x) * a) * of_real x powr a" by (simp add: powr_def exp_add algebra_simps Ln_of_real cis_conv_exp) also note cis_pi finally show ?thesis by simp qed simp_all lemma has_field_derivative_powr: fixes z :: complex assumes "z \ \\<^sub>\\<^sub>0" shows "((\z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)" proof (cases "z=0") case False show ?thesis unfolding powr_def proof (rule has_field_derivative_transform_within) show "((\z. exp (s * Ln z)) has_field_derivative s * (if z = 0 then 0 else exp ((s - 1) * Ln z))) (at z)" apply (intro derivative_eq_intros | simp add: assms)+ by (simp add: False divide_complex_def exp_diff left_diff_distrib') qed (use False in auto) qed (use assms in auto) declare has_field_derivative_powr[THEN DERIV_chain2, derivative_intros] lemma has_field_derivative_powr_of_int: fixes z :: complex assumes gderiv:"(g has_field_derivative gd) (at z within s)" and "g z\0" shows "((\z. g z powr of_int n) has_field_derivative (n * g z powr (of_int n - 1) * gd)) (at z within s)" proof - define dd where "dd = of_int n * g z powr (of_int (n - 1)) * gd" obtain e where "e>0" and e_dist:"\y\s. dist z y < e \ g y \ 0" using DERIV_continuous[OF gderiv,THEN continuous_within_avoid] \g z\0\ by auto have ?thesis when "n\0" proof - define dd' where "dd' = of_int n * g z ^ (nat n - 1) * gd" have "dd=dd'" proof (cases "n=0") case False then have "n-1 \0" using \n\0\ by auto then have "g z powr (of_int (n - 1)) = g z ^ (nat n - 1)" using powr_of_int[OF \g z\0\,of "n-1"] by (simp add: nat_diff_distrib') then show ?thesis unfolding dd_def dd'_def by simp qed (simp add:dd_def dd'_def) then have "((\z. g z powr of_int n) has_field_derivative dd) (at z within s) \ ((\z. g z powr of_int n) has_field_derivative dd') (at z within s)" by simp also have "... \ ((\z. g z ^ nat n) has_field_derivative dd') (at z within s)" proof (rule has_field_derivative_cong_eventually) show "\\<^sub>F x in at z within s. g x powr of_int n = g x ^ nat n" unfolding eventually_at apply (rule exI[where x=e]) using powr_of_int that \e>0\ e_dist by (simp add: dist_commute) qed (use powr_of_int \g z\0\ that in simp) also have "..." unfolding dd'_def using gderiv that by (auto intro!: derivative_eq_intros) finally have "((\z. g z powr of_int n) has_field_derivative dd) (at z within s)" . then show ?thesis unfolding dd_def by simp qed moreover have ?thesis when "n<0" proof - define dd' where "dd' = of_int n / g z ^ (nat (1 - n)) * gd" have "dd=dd'" proof - have "g z powr of_int (n - 1) = inverse (g z ^ nat (1-n))" using powr_of_int[OF \g z\0\,of "n-1"] that by auto then show ?thesis unfolding dd_def dd'_def by (simp add: divide_inverse) qed then have "((\z. g z powr of_int n) has_field_derivative dd) (at z within s) \ ((\z. g z powr of_int n) has_field_derivative dd') (at z within s)" by simp also have "... \ ((\z. inverse (g z ^ nat (-n))) has_field_derivative dd') (at z within s)" proof (rule has_field_derivative_cong_eventually) show "\\<^sub>F x in at z within s. g x powr of_int n = inverse (g x ^ nat (- n))" unfolding eventually_at apply (rule exI[where x=e]) using powr_of_int that \e>0\ e_dist by (simp add: dist_commute) qed (use powr_of_int \g z\0\ that in simp) also have "..." proof - have "nat (- n) + nat (1 - n) - Suc 0 = nat (- n) + nat (- n)" by auto then show ?thesis unfolding dd'_def using gderiv that \g z\0\ by (auto intro!: derivative_eq_intros simp add:field_split_simps power_add[symmetric]) qed finally have "((\z. g z powr of_int n) has_field_derivative dd) (at z within s)" . then show ?thesis unfolding dd_def by simp qed ultimately show ?thesis by force qed lemma field_differentiable_powr_of_int: fixes z :: complex assumes gderiv:"g field_differentiable (at z within s)" and "g z\0" shows "(\z. g z powr of_int n) field_differentiable (at z within s)" using has_field_derivative_powr_of_int assms(2) field_differentiable_def gderiv by blast lemma holomorphic_on_powr_of_int [holomorphic_intros]: assumes "f holomorphic_on s" "\z\s. f z\0" shows "(\z. (f z) powr of_int n) holomorphic_on s" proof (cases "n\0") case True then have "?thesis \ (\z. (f z) ^ nat n) holomorphic_on s" apply (rule_tac holomorphic_cong) using assms(2) by (auto simp add:powr_of_int) moreover have "(\z. (f z) ^ nat n) holomorphic_on s" using assms(1) by (auto intro:holomorphic_intros) ultimately show ?thesis by auto next case False then have "?thesis \ (\z. inverse (f z) ^ nat (-n)) holomorphic_on s" apply (rule_tac holomorphic_cong) using assms(2) by (auto simp add:powr_of_int power_inverse) moreover have "(\z. inverse (f z) ^ nat (-n)) holomorphic_on s" using assms by (auto intro!:holomorphic_intros) ultimately show ?thesis by auto qed lemma has_field_derivative_powr_right [derivative_intros]: "w \ 0 \ ((\z. w powr z) has_field_derivative Ln w * w powr z) (at z)" unfolding powr_def by (intro derivative_eq_intros | simp)+ lemma field_differentiable_powr_right [derivative_intros]: fixes w::complex shows "w \ 0 \ (\z. w powr z) field_differentiable (at z)" using field_differentiable_def has_field_derivative_powr_right by blast lemma holomorphic_on_powr_right [holomorphic_intros]: assumes "f holomorphic_on s" shows "(\z. w powr (f z)) holomorphic_on s" proof (cases "w = 0") case False with assms show ?thesis unfolding holomorphic_on_def field_differentiable_def by (metis (full_types) DERIV_chain' has_field_derivative_powr_right) qed simp lemma holomorphic_on_divide_gen [holomorphic_intros]: assumes f: "f holomorphic_on s" and g: "g holomorphic_on s" and 0: "\z z'. \z \ s; z' \ s\ \ g z = 0 \ g z' = 0" shows "(\z. f z / g z) holomorphic_on s" proof (cases "\z\s. g z = 0") case True with 0 have "g z = 0" if "z \ s" for z using that by blast then show ?thesis using g holomorphic_transform by auto next case False with 0 have "g z \ 0" if "z \ s" for z using that by blast with holomorphic_on_divide show ?thesis using f g by blast qed lemma norm_powr_real_powr: "w \ \ \ 0 \ Re w \ cmod (w powr z) = Re w powr Re z" by (metis dual_order.order_iff_strict norm_powr_real norm_zero of_real_0 of_real_Re powr_def) lemma tendsto_powr_complex: fixes f g :: "_ \ complex" assumes a: "a \ \\<^sub>\\<^sub>0" assumes f: "(f \ a) F" and g: "(g \ b) F" shows "((\z. f z powr g z) \ a powr b) F" proof - from a have [simp]: "a \ 0" by auto from f g a have "((\z. exp (g z * ln (f z))) \ a powr b) F" (is ?P) by (auto intro!: tendsto_intros simp: powr_def) also { have "eventually (\z. z \ 0) (nhds a)" by (intro t1_space_nhds) simp_all with f have "eventually (\z. f z \ 0) F" using filterlim_iff by blast } hence "?P \ ((\z. f z powr g z) \ a powr b) F" by (intro tendsto_cong refl) (simp_all add: powr_def mult_ac) finally show ?thesis . qed lemma tendsto_powr_complex_0: fixes f g :: "'a \ complex" assumes f: "(f \ 0) F" and g: "(g \ b) F" and b: "Re b > 0" shows "((\z. f z powr g z) \ 0) F" proof (rule tendsto_norm_zero_cancel) define h where "h = (\z. if f z = 0 then 0 else exp (Re (g z) * ln (cmod (f z)) + abs (Im (g z)) * pi))" { fix z :: 'a assume z: "f z \ 0" define c where "c = abs (Im (g z)) * pi" from mpi_less_Im_Ln[OF z] Im_Ln_le_pi[OF z] have "abs (Im (Ln (f z))) \ pi" by simp from mult_left_mono[OF this, of "abs (Im (g z))"] have "abs (Im (g z) * Im (ln (f z))) \ c" by (simp add: abs_mult c_def) hence "-Im (g z) * Im (ln (f z)) \ c" by simp hence "norm (f z powr g z) \ h z" by (simp add: powr_def field_simps h_def c_def) } hence le: "norm (f z powr g z) \ h z" for z by (cases "f x = 0") (simp_all add: h_def) have g': "(g \ b) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ g]) simp_all have "((\x. norm (f x)) \ 0) (inf F (principal {z. f z \ 0}))" by (subst tendsto_norm_zero_iff, rule tendsto_mono[OF _ f]) simp_all moreover { have "filterlim (\x. norm (f x)) (principal {0<..}) (principal {z. f z \ 0})" by (auto simp: filterlim_def) hence "filterlim (\x. norm (f x)) (principal {0<..}) (inf F (principal {z. f z \ 0}))" by (rule filterlim_mono) simp_all } ultimately have norm: "filterlim (\x. norm (f x)) (at_right 0) (inf F (principal {z. f z \ 0}))" by (simp add: filterlim_inf at_within_def) have A: "LIM x inf F (principal {z. f z \ 0}). Re (g x) * -ln (cmod (f x)) :> at_top" by (rule filterlim_tendsto_pos_mult_at_top tendsto_intros g' b filterlim_compose[OF filterlim_uminus_at_top_at_bot] filterlim_compose[OF ln_at_0] norm)+ have B: "LIM x inf F (principal {z. f z \ 0}). -\Im (g x)\ * pi + -(Re (g x) * ln (cmod (f x))) :> at_top" by (rule filterlim_tendsto_add_at_top tendsto_intros g')+ (insert A, simp_all) have C: "(h \ 0) F" unfolding h_def by (intro filterlim_If tendsto_const filterlim_compose[OF exp_at_bot]) (insert B, auto simp: filterlim_uminus_at_bot algebra_simps) show "((\x. norm (f x powr g x)) \ 0) F" by (rule Lim_null_comparison[OF always_eventually C]) (insert le, auto) qed lemma tendsto_powr_complex' [tendsto_intros]: fixes f g :: "_ \ complex" assumes "a \ \\<^sub>\\<^sub>0 \ (a = 0 \ Re b > 0)" and "(f \ a) F" "(g \ b) F" shows "((\z. f z powr g z) \ a powr b) F" using assms tendsto_powr_complex tendsto_powr_complex_0 by fastforce lemma tendsto_neg_powr_complex_of_real: assumes "filterlim f at_top F" and "Re s < 0" shows "((\x. complex_of_real (f x) powr s) \ 0) F" proof - have "((\x. norm (complex_of_real (f x) powr s)) \ 0) F" proof (rule Lim_transform_eventually) from assms(1) have "eventually (\x. f x \ 0) F" by (auto simp: filterlim_at_top) thus "eventually (\x. f x powr Re s = norm (of_real (f x) powr s)) F" by eventually_elim (simp add: norm_powr_real_powr) from assms show "((\x. f x powr Re s) \ 0) F" by (intro tendsto_neg_powr) qed thus ?thesis by (simp add: tendsto_norm_zero_iff) qed lemma tendsto_neg_powr_complex_of_nat: assumes "filterlim f at_top F" and "Re s < 0" shows "((\x. of_nat (f x) powr s) \ 0) F" proof - have "((\x. of_real (real (f x)) powr s) \ 0) F" using assms(2) by (intro filterlim_compose[OF _ tendsto_neg_powr_complex_of_real] filterlim_compose[OF _ assms(1)] filterlim_real_sequentially filterlim_ident) auto thus ?thesis by simp qed lemma continuous_powr_complex: assumes "f (netlimit F) \ \\<^sub>\\<^sub>0" "continuous F f" "continuous F g" shows "continuous F (\z. f z powr g z :: complex)" using assms unfolding continuous_def by (intro tendsto_powr_complex) simp_all lemma isCont_powr_complex [continuous_intros]: assumes "f z \ \\<^sub>\\<^sub>0" "isCont f z" "isCont g z" shows "isCont (\z. f z powr g z :: complex) z" using assms unfolding isCont_def by (intro tendsto_powr_complex) simp_all lemma continuous_on_powr_complex [continuous_intros]: assumes "A \ {z. Re (f z) \ 0 \ Im (f z) \ 0}" assumes "\z. z \ A \ f z = 0 \ Re (g z) > 0" assumes "continuous_on A f" "continuous_on A g" shows "continuous_on A (\z. f z powr g z)" unfolding continuous_on_def proof fix z assume z: "z \ A" show "((\z. f z powr g z) \ f z powr g z) (at z within A)" proof (cases "f z = 0") case False from assms(1,2) z have "Re (f z) \ 0 \ Im (f z) \ 0" "f z = 0 \ Re (g z) > 0" by auto with assms(3,4) z show ?thesis by (intro tendsto_powr_complex') (auto elim!: nonpos_Reals_cases simp: complex_eq_iff continuous_on_def) next case True with assms z show ?thesis by (auto intro!: tendsto_powr_complex_0 simp: continuous_on_def) qed qed subsection\<^marker>\tag unimportant\\Some Limits involving Logarithms\ lemma lim_Ln_over_power: fixes s::complex assumes "0 < Re s" shows "(\n. Ln (of_nat n) / of_nat n powr s) \ 0" proof (simp add: lim_sequentially dist_norm, clarify) fix e::real assume e: "0 < e" have "\xo>0. \x\xo. 0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" proof (rule_tac x="2/(e * (Re s)\<^sup>2)" in exI, safe) show "0 < 2 / (e * (Re s)\<^sup>2)" using e assms by (simp add: field_simps) next fix x::real assume x: "2 / (e * (Re s)\<^sup>2) \ x" have "2 / (e * (Re s)\<^sup>2) > 0" using e assms by simp with x have "x > 0" by linarith then have "x * 2 \ e * (x\<^sup>2 * (Re s)\<^sup>2)" using e assms x by (auto simp: power2_eq_square field_simps) also have "... < e * (2 + (x * (Re s * 2) + x\<^sup>2 * (Re s)\<^sup>2))" using e assms \x > 0\ by (auto simp: power2_eq_square field_simps add_pos_pos) finally show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" by (auto simp: algebra_simps) qed then have "\xo>0. \x\xo. x / e < 1 + (Re s * x) + (1/2) * (Re s * x)^2" using e by (simp add: field_simps) then have "\xo>0. \x\xo. x / e < exp (Re s * x)" using assms by (force intro: less_le_trans [OF _ exp_lower_Taylor_quadratic]) then obtain xo where "xo > 0" and xo: "\x. x \ xo \ x < e * exp (Re s * x)" using e by (auto simp: field_simps) have "norm (Ln (of_nat n) / of_nat n powr s) < e" if "n \ nat \exp xo\" for n using e xo [of "ln n"] that apply (auto simp: norm_divide norm_powr_real field_split_simps) apply (metis exp_less_mono exp_ln not_le of_nat_0_less_iff) done then show "\no. \n\no. norm (Ln (of_nat n) / of_nat n powr s) < e" by blast qed lemma lim_Ln_over_n: "((\n. Ln(of_nat n) / of_nat n) \ 0) sequentially" using lim_Ln_over_power [of 1] by simp lemma lim_ln_over_power: fixes s :: real assumes "0 < s" shows "((\n. ln n / (n powr s)) \ 0) sequentially" using lim_Ln_over_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide) done lemma lim_ln_over_n [tendsto_intros]: "((\n. ln(real_of_nat n) / of_nat n) \ 0) sequentially" using lim_ln_over_power [of 1, THEN filterlim_sequentially_Suc [THEN iffD2]] apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm) done lemma lim_log_over_n [tendsto_intros]: "(\n. log k n/n) \ 0" proof - have *: "log k n/n = (1/ln k) * (ln n / n)" for n unfolding log_def by auto have "(\n. (1/ln k) * (ln n / n)) \ (1/ln k) * 0" by (intro tendsto_intros) then show ?thesis unfolding * by auto qed lemma lim_1_over_complex_power: assumes "0 < Re s" shows "(\n. 1 / of_nat n powr s) \ 0" proof (rule Lim_null_comparison) have "\n>0. 3 \ n \ 1 \ ln (real_of_nat n)" using ln_272_gt_1 by (force intro: order_trans [of _ "ln (272/100)"]) then show "\\<^sub>F x in sequentially. cmod (1 / of_nat x powr s) \ cmod (Ln (of_nat x) / of_nat x powr s)" by (auto simp: norm_divide field_split_simps eventually_sequentially) show "(\n. cmod (Ln (of_nat n) / of_nat n powr s)) \ 0" using lim_Ln_over_power [OF assms] by (metis tendsto_norm_zero_iff) qed lemma lim_1_over_real_power: fixes s :: real assumes "0 < s" shows "((\n. 1 / (of_nat n powr s)) \ 0) sequentially" using lim_1_over_complex_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm) apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide) done lemma lim_1_over_Ln: "((\n. 1 / Ln(of_nat n)) \ 0) sequentially" proof (clarsimp simp add: lim_sequentially dist_norm norm_divide field_split_simps) fix r::real assume "0 < r" have ir: "inverse (exp (inverse r)) > 0" by simp obtain n where n: "1 < of_nat n * inverse (exp (inverse r))" using ex_less_of_nat_mult [of _ 1, OF ir] by auto then have "exp (inverse r) < of_nat n" by (simp add: field_split_simps) then have "ln (exp (inverse r)) < ln (of_nat n)" by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff) with \0 < r\ have "1 < r * ln (real_of_nat n)" by (simp add: field_simps) moreover have "n > 0" using n using neq0_conv by fastforce ultimately show "\no. \k. Ln (of_nat k) \ 0 \ no \ k \ 1 < r * cmod (Ln (of_nat k))" using n \0 < r\ by (rule_tac x=n in exI) (force simp: field_split_simps intro: less_le_trans) qed lemma lim_1_over_ln: "((\n. 1 / ln(real_of_nat n)) \ 0) sequentially" using lim_1_over_Ln [THEN filterlim_sequentially_Suc [THEN iffD2]] apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm) apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide) done lemma lim_ln1_over_ln: "(\n. ln(Suc n) / ln n) \ 1" proof (rule Lim_transform_eventually) have "(\n. ln(1 + 1/n) / ln n) \ 0" proof (rule Lim_transform_bound) show "(inverse o real) \ 0" by (metis comp_def lim_inverse_n lim_explicit) show "\\<^sub>F n in sequentially. norm (ln (1 + 1 / n) / ln n) \ norm ((inverse \ real) n)" proof fix n::nat assume n: "3 \ n" then have "ln 3 \ ln n" and ln0: "0 \ ln n" by auto with ln3_gt_1 have "1/ ln n \ 1" by (simp add: field_split_simps) moreover have "ln (1 + 1 / real n) \ 1/n" by (simp add: ln_add_one_self_le_self) ultimately have "ln (1 + 1 / real n) * (1 / ln n) \ (1/n) * 1" by (intro mult_mono) (use n in auto) then show "norm (ln (1 + 1 / n) / ln n) \ norm ((inverse \ real) n)" by (simp add: field_simps ln0) qed qed then show "(\n. 1 + ln(1 + 1/n) / ln n) \ 1" by (metis (full_types) add.right_neutral tendsto_add_const_iff) show "\\<^sub>F k in sequentially. 1 + ln (1 + 1 / k) / ln k = ln(Suc k) / ln k" by (simp add: field_split_simps ln_div eventually_sequentiallyI [of 2]) qed lemma lim_ln_over_ln1: "(\n. ln n / ln(Suc n)) \ 1" proof - have "(\n. inverse (ln(Suc n) / ln n)) \ inverse 1" by (rule tendsto_inverse [OF lim_ln1_over_ln]) auto then show ?thesis by simp qed subsection\<^marker>\tag unimportant\\Relation between Square Root and exp/ln, hence its derivative\ lemma csqrt_exp_Ln: assumes "z \ 0" shows "csqrt z = exp(Ln(z) / 2)" proof - have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))" by (metis exp_double nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) also have "... = z" using assms exp_Ln by blast finally have "csqrt z = csqrt ((exp (Ln z / 2))\<^sup>2)" by simp also have "... = exp (Ln z / 2)" apply (subst csqrt_square) using cos_gt_zero_pi [of "(Im (Ln z) / 2)"] Im_Ln_le_pi mpi_less_Im_Ln assms apply (auto simp: Re_exp Im_exp zero_less_mult_iff zero_le_mult_iff, fastforce+) done finally show ?thesis using assms csqrt_square by simp qed lemma csqrt_inverse: assumes "z \ \\<^sub>\\<^sub>0" shows "csqrt (inverse z) = inverse (csqrt z)" proof (cases "z=0") case False then show ?thesis using assms csqrt_exp_Ln Ln_inverse exp_minus by (simp add: csqrt_exp_Ln Ln_inverse exp_minus) qed auto lemma cnj_csqrt: assumes "z \ \\<^sub>\\<^sub>0" shows "cnj(csqrt z) = csqrt(cnj z)" proof (cases "z=0") case False then show ?thesis by (simp add: assms cnj_Ln csqrt_exp_Ln exp_cnj) qed auto lemma has_field_derivative_csqrt: assumes "z \ \\<^sub>\\<^sub>0" shows "(csqrt has_field_derivative inverse(2 * csqrt z)) (at z)" proof - have z: "z \ 0" using assms by auto then have *: "inverse z = inverse (2*z) * 2" by (simp add: field_split_simps) have [simp]: "exp (Ln z / 2) * inverse z = inverse (csqrt z)" by (simp add: z field_simps csqrt_exp_Ln [symmetric]) (metis power2_csqrt power2_eq_square) have "Im z = 0 \ 0 < Re z" using assms complex_nonpos_Reals_iff not_less by blast with z have "((\z. exp (Ln z / 2)) has_field_derivative inverse (2 * csqrt z)) (at z)" by (force intro: derivative_eq_intros * simp add: assms) then show ?thesis proof (rule has_field_derivative_transform_within) show "\x. dist x z < cmod z \ exp (Ln x / 2) = csqrt x" by (metis csqrt_exp_Ln dist_0_norm less_irrefl) qed (use z in auto) qed lemma field_differentiable_at_csqrt: "z \ \\<^sub>\\<^sub>0 \ csqrt field_differentiable at z" using field_differentiable_def has_field_derivative_csqrt by blast lemma field_differentiable_within_csqrt: "z \ \\<^sub>\\<^sub>0 \ csqrt field_differentiable (at z within s)" using field_differentiable_at_csqrt field_differentiable_within_subset by blast lemma continuous_at_csqrt: "z \ \\<^sub>\\<^sub>0 \ continuous (at z) csqrt" by (simp add: field_differentiable_within_csqrt field_differentiable_imp_continuous_at) corollary\<^marker>\tag unimportant\ isCont_csqrt' [simp]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. csqrt (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_csqrt]) lemma continuous_within_csqrt: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within s) csqrt" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_csqrt) lemma continuous_on_csqrt [continuous_intros]: "(\z. z \ s \ z \ \\<^sub>\\<^sub>0) \ continuous_on s csqrt" by (simp add: continuous_at_imp_continuous_on continuous_within_csqrt) lemma holomorphic_on_csqrt: "(\z. z \ s \ z \ \\<^sub>\\<^sub>0) \ csqrt holomorphic_on s" by (simp add: field_differentiable_within_csqrt holomorphic_on_def) lemma continuous_within_closed_nontrivial: "closed s \ a \ s ==> continuous (at a within s) f" using open_Compl by (force simp add: continuous_def eventually_at_topological filterlim_iff open_Collect_neg) lemma continuous_within_csqrt_posreal: "continuous (at z within (\ \ {w. 0 \ Re(w)})) csqrt" proof (cases "z \ \\<^sub>\\<^sub>0") case True have *: "\e. \0 < e\ \ \x'\\ \ {w. 0 \ Re w}. cmod x' < e^2 \ cmod (csqrt x') < e" by (auto simp: Reals_def real_less_lsqrt) have "Im z = 0" "Re z < 0 \ z = 0" using True cnj.code complex_cnj_zero_iff by (auto simp: Complex_eq complex_nonpos_Reals_iff) fastforce with * show ?thesis apply (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge]) apply (auto simp: continuous_within_eps_delta) using zero_less_power by blast next case False then show ?thesis by (blast intro: continuous_within_csqrt) qed subsection\Complex arctangent\ text\The branch cut gives standard bounds in the real case.\ definition\<^marker>\tag important\ Arctan :: "complex \ complex" where "Arctan \ \z. (\/2) * Ln((1 - \*z) / (1 + \*z))" lemma Arctan_def_moebius: "Arctan z = \/2 * Ln (moebius (-\) 1 \ 1 z)" by (simp add: Arctan_def moebius_def add_ac) lemma Ln_conv_Arctan: assumes "z \ -1" shows "Ln z = -2*\ * Arctan (moebius 1 (- 1) (- \) (- \) z)" proof - have "Arctan (moebius 1 (- 1) (- \) (- \) z) = \/2 * Ln (moebius (- \) 1 \ 1 (moebius 1 (- 1) (- \) (- \) z))" by (simp add: Arctan_def_moebius) also from assms have "\ * z \ \ * (-1)" by (subst mult_left_cancel) simp hence "\ * z - -\ \ 0" by (simp add: eq_neg_iff_add_eq_0) from moebius_inverse'[OF _ this, of 1 1] have "moebius (- \) 1 \ 1 (moebius 1 (- 1) (- \) (- \) z) = z" by simp finally show ?thesis by (simp add: field_simps) qed lemma Arctan_0 [simp]: "Arctan 0 = 0" by (simp add: Arctan_def) lemma Im_complex_div_lemma: "Im((1 - \*z) / (1 + \*z)) = 0 \ Re z = 0" by (auto simp: Im_complex_div_eq_0 algebra_simps) lemma Re_complex_div_lemma: "0 < Re((1 - \*z) / (1 + \*z)) \ norm z < 1" by (simp add: Re_complex_div_gt_0 algebra_simps cmod_def power2_eq_square) lemma tan_Arctan: assumes "z\<^sup>2 \ -1" shows [simp]:"tan(Arctan z) = z" proof - have "1 + \*z \ 0" by (metis assms complex_i_mult_minus i_squared minus_unique power2_eq_square power2_minus) moreover have "1 - \*z \ 0" by (metis assms complex_i_mult_minus i_squared power2_eq_square power2_minus right_minus_eq) ultimately show ?thesis by (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus csqrt_exp_Ln [symmetric] divide_simps power2_eq_square [symmetric]) qed lemma Arctan_tan [simp]: assumes "\Re z\ < pi/2" shows "Arctan(tan z) = z" proof - have ge_pi2: "\n::int. \of_int (2*n + 1) * pi/2\ \ pi/2" by (case_tac n rule: int_cases) (auto simp: abs_mult) have "exp (\*z)*exp (\*z) = -1 \ exp (2*\*z) = -1" by (metis distrib_right exp_add mult_2) also have "... \ exp (2*\*z) = exp (\*pi)" using cis_conv_exp cis_pi by auto also have "... \ exp (2*\*z - \*pi) = 1" by (metis (no_types) diff_add_cancel diff_minus_eq_add exp_add exp_minus_inverse mult.commute) also have "... \ Re(\*2*z - \*pi) = 0 \ (\n::int. Im(\*2*z - \*pi) = of_int (2 * n) * pi)" by (simp add: exp_eq_1) also have "... \ Im z = 0 \ (\n::int. 2 * Re z = of_int (2*n + 1) * pi)" by (simp add: algebra_simps) also have "... \ False" using assms ge_pi2 apply (auto simp: algebra_simps) by (metis abs_mult_pos not_less of_nat_less_0_iff of_nat_numeral) finally have *: "exp (\*z)*exp (\*z) + 1 \ 0" by (auto simp: add.commute minus_unique) show ?thesis using assms * apply (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps i_times_eq_iff power2_eq_square [symmetric]) apply (rule Ln_unique) apply (auto simp: divide_simps exp_minus) apply (simp add: algebra_simps exp_double [symmetric]) done qed lemma assumes "Re z = 0 \ \Im z\ < 1" shows Re_Arctan_bounds: "\Re(Arctan z)\ < pi/2" and has_field_derivative_Arctan: "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" proof - have nz0: "1 + \*z \ 0" using assms by (metis abs_one add_diff_cancel_left' complex_i_mult_minus diff_0 i_squared imaginary_unit.simps less_asym neg_equal_iff_equal) have "z \ -\" using assms by auto then have zz: "1 + z * z \ 0" by (metis abs_one assms i_squared imaginary_unit.simps less_irrefl minus_unique square_eq_iff) have nz1: "1 - \*z \ 0" using assms by (force simp add: i_times_eq_iff) have nz2: "inverse (1 + \*z) \ 0" using assms by (metis Im_complex_div_lemma Re_complex_div_lemma cmod_eq_Im divide_complex_def less_irrefl mult_zero_right zero_complex.simps(1) zero_complex.simps(2)) have nzi: "((1 - \*z) * inverse (1 + \*z)) \ 0" using nz1 nz2 by auto have "Im ((1 - \*z) / (1 + \*z)) = 0 \ 0 < Re ((1 - \*z) / (1 + \*z))" apply (simp add: divide_complex_def) apply (simp add: divide_simps split: if_split_asm) using assms apply (auto simp: algebra_simps abs_square_less_1 [unfolded power2_eq_square]) done then have *: "((1 - \*z) / (1 + \*z)) \ \\<^sub>\\<^sub>0" by (auto simp add: complex_nonpos_Reals_iff) show "\Re(Arctan z)\ < pi/2" unfolding Arctan_def divide_complex_def using mpi_less_Im_Ln [OF nzi] apply (auto simp: abs_if intro!: Im_Ln_less_pi * [unfolded divide_complex_def]) done show "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" unfolding Arctan_def scaleR_conv_of_real apply (intro derivative_eq_intros | simp add: nz0 *)+ using nz0 nz1 zz apply (simp add: algebra_simps field_split_simps power2_eq_square) apply algebra done qed lemma field_differentiable_at_Arctan: "(Re z = 0 \ \Im z\ < 1) \ Arctan field_differentiable at z" using has_field_derivative_Arctan by (auto simp: field_differentiable_def) lemma field_differentiable_within_Arctan: "(Re z = 0 \ \Im z\ < 1) \ Arctan field_differentiable (at z within s)" using field_differentiable_at_Arctan field_differentiable_at_within by blast declare has_field_derivative_Arctan [derivative_intros] declare has_field_derivative_Arctan [THEN DERIV_chain2, derivative_intros] lemma continuous_at_Arctan: "(Re z = 0 \ \Im z\ < 1) \ continuous (at z) Arctan" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Arctan) lemma continuous_within_Arctan: "(Re z = 0 \ \Im z\ < 1) \ continuous (at z within s) Arctan" using continuous_at_Arctan continuous_at_imp_continuous_within by blast lemma continuous_on_Arctan [continuous_intros]: "(\z. z \ s \ Re z = 0 \ \Im z\ < 1) \ continuous_on s Arctan" by (auto simp: continuous_at_imp_continuous_on continuous_within_Arctan) lemma holomorphic_on_Arctan: "(\z. z \ s \ Re z = 0 \ \Im z\ < 1) \ Arctan holomorphic_on s" by (simp add: field_differentiable_within_Arctan holomorphic_on_def) theorem Arctan_series: assumes z: "norm (z :: complex) < 1" defines "g \ \n. if odd n then -\*\^n / n else 0" defines "h \ \z n. (-1)^n / of_nat (2*n+1) * (z::complex)^(2*n+1)" shows "(\n. g n * z^n) sums Arctan z" and "h z sums Arctan z" proof - define G where [abs_def]: "G z = (\n. g n * z^n)" for z have summable: "summable (\n. g n * u^n)" if "norm u < 1" for u proof (cases "u = 0") assume u: "u \ 0" have "(\n. ereal (norm (h u n) / norm (h u (Suc n)))) = (\n. ereal (inverse (norm u)^2) * ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n)))))" proof fix n have "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * ereal (((2*Suc n+1) / (Suc n)) / ((2*Suc n-1) / (Suc n)))" by (simp add: h_def norm_mult norm_power norm_divide field_split_simps power2_eq_square eval_nat_numeral del: of_nat_add of_nat_Suc) also have "of_nat (2*Suc n+1) / of_nat (Suc n) = (2::real) + inverse (real (Suc n))" by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all? also have "of_nat (2*Suc n-1) / of_nat (Suc n) = (2::real) - inverse (real (Suc n))" by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all? finally show "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n))))" . qed also have "\ \ ereal (inverse (norm u)^2) * ereal ((2 + 0) / (2 - 0))" by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) simp_all finally have "liminf (\n. ereal (cmod (h u n) / cmod (h u (Suc n)))) = inverse (norm u)^2" by (intro lim_imp_Liminf) simp_all moreover from power_strict_mono[OF that, of 2] u have "inverse (norm u)^2 > 1" by (simp add: field_split_simps) ultimately have A: "liminf (\n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp from u have "summable (h u)" by (intro summable_norm_cancel[OF ratio_test_convergence[OF _ A]]) (auto simp: h_def norm_divide norm_mult norm_power simp del: of_nat_Suc intro!: mult_pos_pos divide_pos_pos always_eventually) thus "summable (\n. g n * u^n)" by (subst summable_mono_reindex[of "\n. 2*n+1", symmetric]) (auto simp: power_mult strict_mono_def g_def h_def elim!: oddE) qed (simp add: h_def) have "\c. \u\ball 0 1. Arctan u - G u = c" proof (rule has_field_derivative_zero_constant) fix u :: complex assume "u \ ball 0 1" hence u: "norm u < 1" by (simp add: dist_0_norm) define K where "K = (norm u + 1) / 2" from u and abs_Im_le_cmod[of u] have Im_u: "\Im u\ < 1" by linarith from u have K: "0 \ K" "norm u < K" "K < 1" by (simp_all add: K_def) hence "(G has_field_derivative (\n. diffs g n * u ^ n)) (at u)" unfolding G_def by (intro termdiffs_strong[of _ "of_real K"] summable) simp_all also have "(\n. diffs g n * u^n) = (\n. if even n then (\*u)^n else 0)" by (intro ext) (simp_all del: of_nat_Suc add: g_def diffs_def power_mult_distrib) also have "suminf \ = (\n. (-(u^2))^n)" by (subst suminf_mono_reindex[of "\n. 2*n", symmetric]) (auto elim!: evenE simp: strict_mono_def power_mult power_mult_distrib) also from u have "norm u^2 < 1^2" by (intro power_strict_mono) simp_all hence "(\n. (-(u^2))^n) = inverse (1 + u^2)" by (subst suminf_geometric) (simp_all add: norm_power inverse_eq_divide) finally have "(G has_field_derivative inverse (1 + u\<^sup>2)) (at u)" . from DERIV_diff[OF has_field_derivative_Arctan this] Im_u u show "((\u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)" by (simp_all add: at_within_open[OF _ open_ball]) qed simp_all then obtain c where c: "\u. norm u < 1 \ Arctan u - G u = c" by auto from this[of 0] have "c = 0" by (simp add: G_def g_def) with c z have "Arctan z = G z" by simp with summable[OF z] show "(\n. g n * z^n) sums Arctan z" unfolding G_def by (simp add: sums_iff) thus "h z sums Arctan z" by (subst (asm) sums_mono_reindex[of "\n. 2*n+1", symmetric]) (auto elim!: oddE simp: strict_mono_def power_mult g_def h_def) qed text \A quickly-converging series for the logarithm, based on the arctangent.\ theorem ln_series_quadratic: assumes x: "x > (0::real)" shows "(\n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x" proof - define y :: complex where "y = of_real ((x-1)/(x+1))" from x have x': "complex_of_real x \ of_real (-1)" by (subst of_real_eq_iff) auto from x have "\x - 1\ < \x + 1\" by linarith hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1" by (simp add: norm_divide del: of_real_add of_real_diff) hence "norm (\ * y) < 1" unfolding y_def by (subst norm_mult) simp hence "(\n. (-2*\) * ((-1)^n / of_nat (2*n+1) * (\*y)^(2*n+1))) sums ((-2*\) * Arctan (\*y))" by (intro Arctan_series sums_mult) simp_all also have "(\n. (-2*\) * ((-1)^n / of_nat (2*n+1) * (\*y)^(2*n+1))) = (\n. (-2*\) * ((-1)^n * (\*y*(-y\<^sup>2)^n)/of_nat (2*n+1)))" by (intro ext) (simp_all add: power_mult power_mult_distrib) also have "\ = (\n. 2*y* ((-1) * (-y\<^sup>2))^n/of_nat (2*n+1))" by (intro ext, subst power_mult_distrib) (simp add: algebra_simps power_mult) also have "\ = (\n. 2*y^(2*n+1) / of_nat (2*n+1))" by (subst power_add, subst power_mult) (simp add: mult_ac) also have "\ = (\n. of_real (2*((x-1)/(x+1))^(2*n+1) / of_nat (2*n+1)))" by (intro ext) (simp add: y_def) also have "\ * y = (of_real x - 1) / (-\ * (of_real x + 1))" by (subst divide_divide_eq_left [symmetric]) (simp add: y_def) also have "\ = moebius 1 (-1) (-\) (-\) (of_real x)" by (simp add: moebius_def algebra_simps) also from x' have "-2*\*Arctan \ = Ln (of_real x)" by (intro Ln_conv_Arctan [symmetric]) simp_all also from x have "\ = ln x" by (rule Ln_of_real) finally show ?thesis by (subst (asm) sums_of_real_iff) qed subsection\<^marker>\tag unimportant\ \Real arctangent\ lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0" proof - have ne: "1 + x\<^sup>2 \ 0" by (metis power_one sum_power2_eq_zero_iff zero_neq_one) have "Re (Ln ((1 - \ * x) * inverse (1 + \ * x))) = 0" apply (rule norm_exp_imaginary) apply (subst exp_Ln) using ne apply (simp_all add: cmod_def complex_eq_iff) apply (auto simp: field_split_simps) apply algebra done then show ?thesis unfolding Arctan_def divide_complex_def by (simp add: complex_eq_iff) qed lemma arctan_eq_Re_Arctan: "arctan x = Re (Arctan (of_real x))" proof (rule arctan_unique) show "- (pi / 2) < Re (Arctan (complex_of_real x))" apply (simp add: Arctan_def) apply (rule Im_Ln_less_pi) apply (auto simp: Im_complex_div_lemma complex_nonpos_Reals_iff) done next have *: " (1 - \*x) / (1 + \*x) \ 0" by (simp add: field_split_simps) ( simp add: complex_eq_iff) show "Re (Arctan (complex_of_real x)) < pi / 2" using mpi_less_Im_Ln [OF *] by (simp add: Arctan_def) next have "tan (Re (Arctan (of_real x))) = Re (tan (Arctan (of_real x)))" apply (auto simp: tan_def Complex.Re_divide Re_sin Re_cos Im_sin Im_cos) apply (simp add: field_simps) by (simp add: power2_eq_square) also have "... = x" apply (subst tan_Arctan, auto) by (metis diff_0_right minus_diff_eq mult_zero_left not_le of_real_1 of_real_eq_iff of_real_minus of_real_power power2_eq_square real_minus_mult_self_le zero_less_one) finally show "tan (Re (Arctan (complex_of_real x))) = x" . qed lemma Arctan_of_real: "Arctan (of_real x) = of_real (arctan x)" unfolding arctan_eq_Re_Arctan divide_complex_def by (simp add: complex_eq_iff) lemma Arctan_in_Reals [simp]: "z \ \ \ Arctan z \ \" by (metis Reals_cases Reals_of_real Arctan_of_real) declare arctan_one [simp] lemma arctan_less_pi4_pos: "x < 1 \ arctan x < pi/4" by (metis arctan_less_iff arctan_one) lemma arctan_less_pi4_neg: "-1 < x \ -(pi/4) < arctan x" by (metis arctan_less_iff arctan_minus arctan_one) lemma arctan_less_pi4: "\x\ < 1 \ \arctan x\ < pi/4" by (metis abs_less_iff arctan_less_pi4_pos arctan_minus) lemma arctan_le_pi4: "\x\ \ 1 \ \arctan x\ \ pi/4" by (metis abs_le_iff arctan_le_iff arctan_minus arctan_one) lemma abs_arctan: "\arctan x\ = arctan \x\" by (simp add: abs_if arctan_minus) lemma arctan_add_raw: assumes "\arctan x + arctan y\ < pi/2" shows "arctan x + arctan y = arctan((x + y) / (1 - x * y))" proof (rule arctan_unique [symmetric]) show 12: "- (pi / 2) < arctan x + arctan y" "arctan x + arctan y < pi / 2" using assms by linarith+ show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)" using cos_gt_zero_pi [OF 12] by (simp add: arctan tan_add) qed lemma arctan_inverse: assumes "0 < x" shows "arctan(inverse x) = pi/2 - arctan x" proof - have "arctan(inverse x) = arctan(inverse(tan(arctan x)))" by (simp add: arctan) also have "... = arctan (tan (pi / 2 - arctan x))" by (simp add: tan_cot) also have "... = pi/2 - arctan x" proof - have "0 < pi - arctan x" using arctan_ubound [of x] pi_gt_zero by linarith with assms show ?thesis by (simp add: Transcendental.arctan_tan) qed finally show ?thesis . qed lemma arctan_add_small: assumes "\x * y\ < 1" shows "(arctan x + arctan y = arctan((x + y) / (1 - x * y)))" proof (cases "x = 0 \ y = 0") case True then show ?thesis by auto next case False then have *: "\arctan x\ < pi / 2 - \arctan y\" using assms apply (auto simp add: abs_arctan arctan_inverse [symmetric] arctan_less_iff) apply (simp add: field_split_simps abs_mult) done show ?thesis apply (rule arctan_add_raw) using * by linarith qed lemma abs_arctan_le: fixes x::real shows "\arctan x\ \ \x\" proof - have 1: "\x. x \ \ \ cmod (inverse (1 + x\<^sup>2)) \ 1" by (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square) have "cmod (Arctan w - Arctan z) \ 1 * cmod (w-z)" if "w \ \" "z \ \" for w z apply (rule field_differentiable_bound [OF convex_Reals, of Arctan _ 1]) apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan]) using 1 that apply (auto simp: Reals_def) done then have "cmod (Arctan (of_real x) - Arctan 0) \ 1 * cmod (of_real x -0)" using Reals_0 Reals_of_real by blast then show ?thesis by (simp add: Arctan_of_real) qed lemma arctan_le_self: "0 \ x \ arctan x \ x" by (metis abs_arctan_le abs_of_nonneg zero_le_arctan_iff) lemma abs_tan_ge: "\x\ < pi/2 \ \x\ \ \tan x\" by (metis abs_arctan_le abs_less_iff arctan_tan minus_less_iff) lemma arctan_bounds: assumes "0 \ x" "x < 1" shows arctan_lower_bound: "(\k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \ arctan x" (is "(\k<_. (- 1)^ k * ?a k) \ _") and arctan_upper_bound: "arctan x \ (\k<2 * n + 1. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" proof - have tendsto_zero: "?a \ 0" proof (rule tendsto_eq_rhs) show "(\k. 1 / real (k * 2 + 1) * x ^ (k * 2 + 1)) \ 0 * 0" using assms by (intro tendsto_mult real_tendsto_divide_at_top) (auto simp: filterlim_real_sequentially filterlim_sequentially_iff_filterlim_real intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top) qed simp have nonneg: "0 \ ?a n" for n by (force intro!: divide_nonneg_nonneg mult_nonneg_nonneg zero_le_power assms) have le: "?a (Suc n) \ ?a n" for n by (rule mult_mono[OF _ power_decreasing]) (auto simp: field_split_simps assms less_imp_le) from summable_Leibniz'(4)[of ?a, OF tendsto_zero nonneg le, of n] summable_Leibniz'(2)[of ?a, OF tendsto_zero nonneg le, of n] assms show "(\k<2*n. (- 1)^ k * ?a k) \ arctan x" "arctan x \ (\k<2 * n + 1. (- 1)^ k * ?a k)" by (auto simp: arctan_series) qed subsection\<^marker>\tag unimportant\ \Bounds on pi using real arctangent\ lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)" using machin by simp lemma pi_approx: "3.141592653588 \ pi" "pi \ 3.1415926535899" unfolding pi_machin using arctan_bounds[of "1/5" 4] arctan_bounds[of "1/239" 4] by (simp_all add: eval_nat_numeral) lemma pi_gt3: "pi > 3" using pi_approx by simp subsection\Inverse Sine\ definition\<^marker>\tag important\ Arcsin :: "complex \ complex" where "Arcsin \ \z. -\ * Ln(\ * z + csqrt(1 - z\<^sup>2))" lemma Arcsin_body_lemma: "\ * z + csqrt(1 - z\<^sup>2) \ 0" using power2_csqrt [of "1 - z\<^sup>2"] apply auto by (metis complex_i_mult_minus diff_add_cancel diff_minus_eq_add diff_self mult.assoc mult.left_commute numeral_One power2_csqrt power2_eq_square zero_neq_numeral) lemma Arcsin_range_lemma: "\Re z\ < 1 \ 0 < Re(\ * z + csqrt(1 - z\<^sup>2))" using Complex.cmod_power2 [of z, symmetric] by (simp add: real_less_rsqrt algebra_simps Re_power2 cmod_square_less_1_plus) lemma Re_Arcsin: "Re(Arcsin z) = Im (Ln (\ * z + csqrt(1 - z\<^sup>2)))" by (simp add: Arcsin_def) lemma Im_Arcsin: "Im(Arcsin z) = - ln (cmod (\ * z + csqrt (1 - z\<^sup>2)))" by (simp add: Arcsin_def Arcsin_body_lemma) lemma one_minus_z2_notin_nonpos_Reals: assumes "(Im z = 0 \ \Re z\ < 1)" shows "1 - z\<^sup>2 \ \\<^sub>\\<^sub>0" using assms apply (auto simp: complex_nonpos_Reals_iff Re_power2 Im_power2) using power2_less_0 [of "Im z"] apply force using abs_square_less_1 not_le by blast lemma isCont_Arcsin_lemma: assumes le0: "Re (\ * z + csqrt (1 - z\<^sup>2)) \ 0" and "(Im z = 0 \ \Re z\ < 1)" shows False proof (cases "Im z = 0") case True then show ?thesis using assms by (fastforce simp: cmod_def abs_square_less_1 [symmetric]) next case False have leim: "(cmod (1 - z\<^sup>2) + (1 - Re (z\<^sup>2))) / 2 \ (Im z)\<^sup>2" using le0 sqrt_le_D by fastforce have neq: "(cmod z)\<^sup>2 \ 1 + cmod (1 - z\<^sup>2)" proof (clarsimp simp add: cmod_def) assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 = 1 + sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" by simp then show False using False by (simp add: power2_eq_square algebra_simps) qed moreover have 2: "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2" using leim cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1] by (simp add: norm_power Re_power2 norm_minus_commute [of 1]) ultimately show False by (simp add: Re_power2 Im_power2 cmod_power2) qed lemma isCont_Arcsin: assumes "(Im z = 0 \ \Re z\ < 1)" shows "isCont Arcsin z" proof - have 1: "\ * z + csqrt (1 - z\<^sup>2) \ \\<^sub>\\<^sub>0" by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff) have 2: "1 - z\<^sup>2 \ \\<^sub>\\<^sub>0" by (simp add: one_minus_z2_notin_nonpos_Reals assms) show ?thesis using assms unfolding Arcsin_def by (intro isCont_Ln' isCont_csqrt' continuous_intros 1 2) qed lemma isCont_Arcsin' [simp]: shows "isCont f z \ (Im (f z) = 0 \ \Re (f z)\ < 1) \ isCont (\x. Arcsin (f x)) z" by (blast intro: isCont_o2 [OF _ isCont_Arcsin]) lemma sin_Arcsin [simp]: "sin(Arcsin z) = z" proof - have "\*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \ (\*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0" by (simp add: algebra_simps) \ \Cancelling a factor of 2\ moreover have "... \ (\*z) + csqrt (1 - z\<^sup>2) = 0" by (metis Arcsin_body_lemma distrib_right no_zero_divisors zero_neq_numeral) ultimately show ?thesis apply (simp add: sin_exp_eq Arcsin_def Arcsin_body_lemma exp_minus divide_simps) apply (simp add: algebra_simps) apply (simp add: power2_eq_square [symmetric] algebra_simps) done qed lemma Re_eq_pihalf_lemma: "\Re z\ = pi/2 \ Im z = 0 \ Re ((exp (\*z) + inverse (exp (\*z))) / 2) = 0 \ 0 \ Im ((exp (\*z) + inverse (exp (\*z))) / 2)" apply (simp add: cos_i_times [symmetric] Re_cos Im_cos abs_if del: eq_divide_eq_numeral1) by (metis cos_minus cos_pi_half) lemma Re_less_pihalf_lemma: assumes "\Re z\ < pi / 2" shows "0 < Re ((exp (\*z) + inverse (exp (\*z))) / 2)" proof - have "0 < cos (Re z)" using assms using cos_gt_zero_pi by auto then show ?thesis by (simp add: cos_i_times [symmetric] Re_cos Im_cos add_pos_pos) qed lemma Arcsin_sin: assumes "\Re z\ < pi/2 \ (\Re z\ = pi/2 \ Im z = 0)" shows "Arcsin(sin z) = z" proof - have "Arcsin(sin z) = - (\ * Ln (csqrt (1 - (\ * (exp (\*z) - inverse (exp (\*z))))\<^sup>2 / 4) - (inverse (exp (\*z)) - exp (\*z)) / 2))" by (simp add: sin_exp_eq Arcsin_def exp_minus power_divide) also have "... = - (\ * Ln (csqrt (((exp (\*z) + inverse (exp (\*z)))/2)\<^sup>2) - (inverse (exp (\*z)) - exp (\*z)) / 2))" by (simp add: field_simps power2_eq_square) also have "... = - (\ * Ln (((exp (\*z) + inverse (exp (\*z)))/2) - (inverse (exp (\*z)) - exp (\*z)) / 2))" apply (subst csqrt_square) using assms Re_eq_pihalf_lemma Re_less_pihalf_lemma apply auto done also have "... = - (\ * Ln (exp (\*z)))" by (simp add: field_simps power2_eq_square) also have "... = z" using assms by (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: if_split_asm) finally show ?thesis . qed lemma Arcsin_unique: "\sin z = w; \Re z\ < pi/2 \ (\Re z\ = pi/2 \ Im z = 0)\ \ Arcsin w = z" by (metis Arcsin_sin) lemma Arcsin_0 [simp]: "Arcsin 0 = 0" by (metis Arcsin_sin norm_zero pi_half_gt_zero real_norm_def sin_zero zero_complex.simps(1)) lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2" by (metis Arcsin_sin Im_complex_of_real Re_complex_of_real numeral_One of_real_numeral pi_half_ge_zero real_sqrt_abs real_sqrt_pow2 real_sqrt_power sin_of_real sin_pi_half) lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)" by (metis Arcsin_1 Arcsin_sin Im_complex_of_real Re_complex_of_real abs_of_nonneg of_real_minus pi_half_ge_zero power2_minus real_sqrt_abs sin_Arcsin sin_minus) lemma has_field_derivative_Arcsin: assumes "Im z = 0 \ \Re z\ < 1" shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)" proof - have "(sin (Arcsin z))\<^sup>2 \ 1" using assms one_minus_z2_notin_nonpos_Reals by force then have "cos (Arcsin z) \ 0" by (metis diff_0_right power_zero_numeral sin_squared_eq) then show ?thesis by (rule has_field_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) (auto intro: isCont_Arcsin assms) qed declare has_field_derivative_Arcsin [derivative_intros] declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ Arcsin field_differentiable at z" using field_differentiable_def has_field_derivative_Arcsin by blast lemma field_differentiable_within_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ Arcsin field_differentiable (at z within s)" using field_differentiable_at_Arcsin field_differentiable_within_subset by blast lemma continuous_within_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ continuous (at z within s) Arcsin" using continuous_at_imp_continuous_within isCont_Arcsin by blast lemma continuous_on_Arcsin [continuous_intros]: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ continuous_on s Arcsin" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_Arcsin: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ Arcsin holomorphic_on s" by (simp add: field_differentiable_within_Arcsin holomorphic_on_def) subsection\Inverse Cosine\ definition\<^marker>\tag important\ Arccos :: "complex \ complex" where "Arccos \ \z. -\ * Ln(z + \ * csqrt(1 - z\<^sup>2))" lemma Arccos_range_lemma: "\Re z\ < 1 \ 0 < Im(z + \ * csqrt(1 - z\<^sup>2))" using Arcsin_range_lemma [of "-z"] by simp lemma Arccos_body_lemma: "z + \ * csqrt(1 - z\<^sup>2) \ 0" using Arcsin_body_lemma [of z] by (metis Arcsin_body_lemma complex_i_mult_minus diff_minus_eq_add power2_minus right_minus_eq) lemma Re_Arccos: "Re(Arccos z) = Im (Ln (z + \ * csqrt(1 - z\<^sup>2)))" by (simp add: Arccos_def) lemma Im_Arccos: "Im(Arccos z) = - ln (cmod (z + \ * csqrt (1 - z\<^sup>2)))" by (simp add: Arccos_def Arccos_body_lemma) text\A very tricky argument to find!\ lemma isCont_Arccos_lemma: assumes eq0: "Im (z + \ * csqrt (1 - z\<^sup>2)) = 0" and "(Im z = 0 \ \Re z\ < 1)" shows False proof (cases "Im z = 0") case True then show ?thesis using assms by (fastforce simp add: cmod_def abs_square_less_1 [symmetric]) next case False have Imz: "Im z = - sqrt ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)" using eq0 abs_Re_le_cmod [of "1-z\<^sup>2"] by (simp add: Re_power2 algebra_simps) have "(cmod z)\<^sup>2 - 1 \ cmod (1 - z\<^sup>2)" proof (clarsimp simp add: cmod_def) assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 - 1 = sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" by simp then show False using False by (simp add: power2_eq_square algebra_simps) qed moreover have "(Im z)\<^sup>2 = ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)" apply (subst Imz) using abs_Re_le_cmod [of "1-z\<^sup>2"] apply (simp add: Re_power2) done ultimately show False by (simp add: cmod_power2) qed lemma isCont_Arccos: assumes "(Im z = 0 \ \Re z\ < 1)" shows "isCont Arccos z" proof - have "z + \ * csqrt (1 - z\<^sup>2) \ \\<^sub>\\<^sub>0" by (metis complex_nonpos_Reals_iff isCont_Arccos_lemma assms) with assms show ?thesis apply (simp add: Arccos_def) apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+ apply (simp_all add: one_minus_z2_notin_nonpos_Reals assms) done qed lemma isCont_Arccos' [simp]: shows "isCont f z \ (Im (f z) = 0 \ \Re (f z)\ < 1) \ isCont (\x. Arccos (f x)) z" by (blast intro: isCont_o2 [OF _ isCont_Arccos]) lemma cos_Arccos [simp]: "cos(Arccos z) = z" proof - have "z*2 + \ * (2 * csqrt (1 - z\<^sup>2)) = 0 \ z*2 + \ * csqrt (1 - z\<^sup>2)*2 = 0" by (simp add: algebra_simps) \ \Cancelling a factor of 2\ moreover have "... \ z + \ * csqrt (1 - z\<^sup>2) = 0" by (metis distrib_right mult_eq_0_iff zero_neq_numeral) ultimately show ?thesis apply (simp add: cos_exp_eq Arccos_def Arccos_body_lemma exp_minus field_simps) apply (simp add: power2_eq_square [symmetric]) done qed lemma Arccos_cos: assumes "0 < Re z & Re z < pi \ Re z = 0 & 0 \ Im z \ Re z = pi & Im z \ 0" shows "Arccos(cos z) = z" proof - have *: "((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z))) = sin z" by (simp add: sin_exp_eq exp_minus field_simps power2_eq_square) have "1 - (exp (\ * z) + inverse (exp (\ * z)))\<^sup>2 / 4 = ((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))\<^sup>2" by (simp add: field_simps power2_eq_square) then have "Arccos(cos z) = - (\ * Ln ((exp (\ * z) + inverse (exp (\ * z))) / 2 + \ * csqrt (((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))\<^sup>2)))" by (simp add: cos_exp_eq Arccos_def exp_minus power_divide) also have "... = - (\ * Ln ((exp (\ * z) + inverse (exp (\ * z))) / 2 + \ * ((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))))" apply (subst csqrt_square) using assms Re_sin_pos [of z] Im_sin_nonneg [of z] Im_sin_nonneg2 [of z] apply (auto simp: * Re_sin Im_sin) done also have "... = - (\ * Ln (exp (\*z)))" by (simp add: field_simps power2_eq_square) also have "... = z" using assms apply (subst Complex_Transcendental.Ln_exp, auto) done finally show ?thesis . qed lemma Arccos_unique: "\cos z = w; 0 < Re z \ Re z < pi \ Re z = 0 \ 0 \ Im z \ Re z = pi \ Im z \ 0\ \ Arccos w = z" using Arccos_cos by blast lemma Arccos_0 [simp]: "Arccos 0 = pi/2" by (rule Arccos_unique) auto lemma Arccos_1 [simp]: "Arccos 1 = 0" by (rule Arccos_unique) auto lemma Arccos_minus1: "Arccos(-1) = pi" by (rule Arccos_unique) auto lemma has_field_derivative_Arccos: assumes "(Im z = 0 \ \Re z\ < 1)" shows "(Arccos has_field_derivative - inverse(sin(Arccos z))) (at z)" proof - have "x\<^sup>2 \ -1" for x::real by (sos "((R<1 + (([~1] * A=0) + (R<1 * (R<1 * [x__]^2)))))") with assms have "(cos (Arccos z))\<^sup>2 \ 1" by (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1) then have "- sin (Arccos z) \ 0" by (metis cos_squared_eq diff_0_right mult_zero_left neg_0_equal_iff_equal power2_eq_square) then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)" by (rule has_field_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]]) (auto intro: isCont_Arccos assms) then show ?thesis by simp qed declare has_field_derivative_Arcsin [derivative_intros] declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Arccos: "(Im z = 0 \ \Re z\ < 1) \ Arccos field_differentiable at z" using field_differentiable_def has_field_derivative_Arccos by blast lemma field_differentiable_within_Arccos: "(Im z = 0 \ \Re z\ < 1) \ Arccos field_differentiable (at z within s)" using field_differentiable_at_Arccos field_differentiable_within_subset by blast lemma continuous_within_Arccos: "(Im z = 0 \ \Re z\ < 1) \ continuous (at z within s) Arccos" using continuous_at_imp_continuous_within isCont_Arccos by blast lemma continuous_on_Arccos [continuous_intros]: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ continuous_on s Arccos" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_Arccos: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ Arccos holomorphic_on s" by (simp add: field_differentiable_within_Arccos holomorphic_on_def) subsection\<^marker>\tag unimportant\\Upper and Lower Bounds for Inverse Sine and Cosine\ lemma Arcsin_bounds: "\Re z\ < 1 \ \Re(Arcsin z)\ < pi/2" unfolding Re_Arcsin by (blast intro: Re_Ln_pos_lt_imp Arcsin_range_lemma) lemma Arccos_bounds: "\Re z\ < 1 \ 0 < Re(Arccos z) \ Re(Arccos z) < pi" unfolding Re_Arccos by (blast intro!: Im_Ln_pos_lt_imp Arccos_range_lemma) lemma Re_Arccos_bounds: "-pi < Re(Arccos z) \ Re(Arccos z) \ pi" unfolding Re_Arccos by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arccos_body_lemma) lemma Re_Arccos_bound: "\Re(Arccos z)\ \ pi" by (meson Re_Arccos_bounds abs_le_iff less_eq_real_def minus_less_iff) lemma Im_Arccos_bound: "\Im (Arccos w)\ \ cmod w" proof - have "(Im (Arccos w))\<^sup>2 \ (cmod (cos (Arccos w)))\<^sup>2 - (cos (Re (Arccos w)))\<^sup>2" using norm_cos_squared [of "Arccos w"] real_le_abs_sinh [of "Im (Arccos w)"] apply (simp only: abs_le_square_iff) apply (simp add: field_split_simps) done also have "... \ (cmod w)\<^sup>2" by (auto simp: cmod_power2) finally show ?thesis using abs_le_square_iff by force qed lemma Re_Arcsin_bounds: "-pi < Re(Arcsin z) & Re(Arcsin z) \ pi" unfolding Re_Arcsin by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arcsin_body_lemma) lemma Re_Arcsin_bound: "\Re(Arcsin z)\ \ pi" by (meson Re_Arcsin_bounds abs_le_iff less_eq_real_def minus_less_iff) lemma norm_Arccos_bounded: fixes w :: complex shows "norm (Arccos w) \ pi + norm w" proof - have Re: "(Re (Arccos w))\<^sup>2 \ pi\<^sup>2" "(Im (Arccos w))\<^sup>2 \ (cmod w)\<^sup>2" using Re_Arccos_bound [of w] Im_Arccos_bound [of w] abs_le_square_iff by force+ have "Arccos w \ Arccos w \ pi\<^sup>2 + (cmod w)\<^sup>2" using Re by (simp add: dot_square_norm cmod_power2 [of "Arccos w"]) then have "cmod (Arccos w) \ pi + cmod (cos (Arccos w))" apply (simp add: norm_le_square) by (metis dot_square_norm norm_ge_zero norm_le_square pi_ge_zero triangle_lemma) then show "cmod (Arccos w) \ pi + cmod w" by auto qed subsection\<^marker>\tag unimportant\\Interrelations between Arcsin and Arccos\ lemma cos_Arcsin_nonzero: assumes "z\<^sup>2 \ 1" shows "cos(Arcsin z) \ 0" proof - have eq: "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = z\<^sup>2 * (z\<^sup>2 - 1)" by (simp add: power_mult_distrib algebra_simps) have "\ * z * (csqrt (1 - z\<^sup>2)) \ z\<^sup>2 - 1" proof assume "\ * z * (csqrt (1 - z\<^sup>2)) = z\<^sup>2 - 1" then have "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (z\<^sup>2 - 1)\<^sup>2" by simp then have "z\<^sup>2 * (z\<^sup>2 - 1) = (z\<^sup>2 - 1)*(z\<^sup>2 - 1)" using eq power2_eq_square by auto then show False using assms by simp qed then have "1 + \ * z * (csqrt (1 - z * z)) \ z\<^sup>2" by (metis add_minus_cancel power2_eq_square uminus_add_conv_diff) then have "2*(1 + \ * z * (csqrt (1 - z * z))) \ 2*z\<^sup>2" (*FIXME cancel_numeral_factor*) by (metis mult_cancel_left zero_neq_numeral) then have "(\ * z + csqrt (1 - z\<^sup>2))\<^sup>2 \ -1" using assms apply (auto simp: power2_sum) apply (simp add: power2_eq_square algebra_simps) done then show ?thesis apply (simp add: cos_exp_eq Arcsin_def exp_minus) apply (simp add: divide_simps Arcsin_body_lemma) apply (metis add.commute minus_unique power2_eq_square) done qed lemma sin_Arccos_nonzero: assumes "z\<^sup>2 \ 1" shows "sin(Arccos z) \ 0" proof - have eq: "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = -(z\<^sup>2) * (1 - z\<^sup>2)" by (simp add: power_mult_distrib algebra_simps) have "\ * z * (csqrt (1 - z\<^sup>2)) \ 1 - z\<^sup>2" proof assume "\ * z * (csqrt (1 - z\<^sup>2)) = 1 - z\<^sup>2" then have "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (1 - z\<^sup>2)\<^sup>2" by simp then have "-(z\<^sup>2) * (1 - z\<^sup>2) = (1 - z\<^sup>2)*(1 - z\<^sup>2)" using eq power2_eq_square by auto then have "-(z\<^sup>2) = (1 - z\<^sup>2)" using assms by (metis add.commute add.right_neutral diff_add_cancel mult_right_cancel) then show False using assms by simp qed then have "z\<^sup>2 + \ * z * (csqrt (1 - z\<^sup>2)) \ 1" by (simp add: algebra_simps) then have "2*(z\<^sup>2 + \ * z * (csqrt (1 - z\<^sup>2))) \ 2*1" by (metis mult_cancel_left2 zero_neq_numeral) (*FIXME cancel_numeral_factor*) then have "(z + \ * csqrt (1 - z\<^sup>2))\<^sup>2 \ 1" using assms apply (auto simp: Power.comm_semiring_1_class.power2_sum power_mult_distrib) apply (simp add: power2_eq_square algebra_simps) done then show ?thesis apply (simp add: sin_exp_eq Arccos_def exp_minus) apply (simp add: divide_simps Arccos_body_lemma) apply (simp add: power2_eq_square) done qed lemma cos_sin_csqrt: assumes "0 < cos(Re z) \ cos(Re z) = 0 \ Im z * sin(Re z) \ 0" shows "cos z = csqrt(1 - (sin z)\<^sup>2)" apply (rule csqrt_unique [THEN sym]) apply (simp add: cos_squared_eq) using assms apply (auto simp: Re_cos Im_cos add_pos_pos mult_le_0_iff zero_le_mult_iff) done lemma sin_cos_csqrt: assumes "0 < sin(Re z) \ sin(Re z) = 0 \ 0 \ Im z * cos(Re z)" shows "sin z = csqrt(1 - (cos z)\<^sup>2)" apply (rule csqrt_unique [THEN sym]) apply (simp add: sin_squared_eq) using assms apply (auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff) done lemma Arcsin_Arccos_csqrt_pos: "(0 < Re z | Re z = 0 & 0 \ Im z) \ Arcsin z = Arccos(csqrt(1 - z\<^sup>2))" by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) lemma Arccos_Arcsin_csqrt_pos: "(0 < Re z | Re z = 0 & 0 \ Im z) \ Arccos z = Arcsin(csqrt(1 - z\<^sup>2))" by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) lemma sin_Arccos: "0 < Re z | Re z = 0 & 0 \ Im z \ sin(Arccos z) = csqrt(1 - z\<^sup>2)" by (simp add: Arccos_Arcsin_csqrt_pos) lemma cos_Arcsin: "0 < Re z | Re z = 0 & 0 \ Im z \ cos(Arcsin z) = csqrt(1 - z\<^sup>2)" by (simp add: Arcsin_Arccos_csqrt_pos) subsection\<^marker>\tag unimportant\\Relationship with Arcsin on the Real Numbers\ lemma Im_Arcsin_of_real: assumes "\x\ \ 1" shows "Im (Arcsin (of_real x)) = 0" proof - have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) then have "cmod (\ * of_real x + csqrt (1 - (of_real x)\<^sup>2))^2 = 1" using assms abs_square_le_1 by (force simp add: Complex.cmod_power2) then have "cmod (\ * of_real x + csqrt (1 - (of_real x)\<^sup>2)) = 1" by (simp add: norm_complex_def) then show ?thesis by (simp add: Im_Arcsin exp_minus) qed corollary\<^marker>\tag unimportant\ Arcsin_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arcsin z \ \" by (metis Im_Arcsin_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) lemma arcsin_eq_Re_Arcsin: assumes "\x\ \ 1" shows "arcsin x = Re (Arcsin (of_real x))" unfolding arcsin_def proof (rule the_equality, safe) show "- (pi / 2) \ Re (Arcsin (complex_of_real x))" using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arcsin) next show "Re (Arcsin (complex_of_real x)) \ pi / 2" using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arcsin) next show "sin (Re (Arcsin (complex_of_real x))) = x" using Re_sin [of "Arcsin (of_real x)"] Arcsin_body_lemma [of "of_real x"] by (simp add: Im_Arcsin_of_real assms) next fix x' assume "- (pi / 2) \ x'" "x' \ pi / 2" "x = sin x'" then show "x' = Re (Arcsin (complex_of_real (sin x')))" apply (simp add: sin_of_real [symmetric]) apply (subst Arcsin_sin) apply (auto simp: ) done qed lemma of_real_arcsin: "\x\ \ 1 \ of_real(arcsin x) = Arcsin(of_real x)" by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0) subsection\<^marker>\tag unimportant\\Relationship with Arccos on the Real Numbers\ lemma Im_Arccos_of_real: assumes "\x\ \ 1" shows "Im (Arccos (of_real x)) = 0" proof - have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) then have "cmod (of_real x + \ * csqrt (1 - (of_real x)\<^sup>2))^2 = 1" using assms abs_square_le_1 by (force simp add: Complex.cmod_power2) then have "cmod (of_real x + \ * csqrt (1 - (of_real x)\<^sup>2)) = 1" by (simp add: norm_complex_def) then show ?thesis by (simp add: Im_Arccos exp_minus) qed corollary\<^marker>\tag unimportant\ Arccos_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arccos z \ \" by (metis Im_Arccos_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) lemma arccos_eq_Re_Arccos: assumes "\x\ \ 1" shows "arccos x = Re (Arccos (of_real x))" unfolding arccos_def proof (rule the_equality, safe) show "0 \ Re (Arccos (complex_of_real x))" using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arccos) next show "Re (Arccos (complex_of_real x)) \ pi" using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arccos) next show "cos (Re (Arccos (complex_of_real x))) = x" using Re_cos [of "Arccos (of_real x)"] Arccos_body_lemma [of "of_real x"] by (simp add: Im_Arccos_of_real assms) next fix x' assume "0 \ x'" "x' \ pi" "x = cos x'" then show "x' = Re (Arccos (complex_of_real (cos x')))" apply (simp add: cos_of_real [symmetric]) apply (subst Arccos_cos) apply (auto simp: ) done qed lemma of_real_arccos: "\x\ \ 1 \ of_real(arccos x) = Arccos(of_real x)" by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0) subsection\<^marker>\tag unimportant\\Some interrelationships among the real inverse trig functions\ lemma arccos_arctan: assumes "-1 < x" "x < 1" shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))" proof - have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0" proof (rule sin_eq_0_pi) show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)" using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x) < pi" using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)) = 0" using assms by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan power2_eq_square square_eq_1_iff) qed then show ?thesis by simp qed lemma arcsin_plus_arccos: assumes "-1 \ x" "x \ 1" shows "arcsin x + arccos x = pi/2" proof - have "arcsin x = pi/2 - arccos x" apply (rule sin_inj_pi) using assms arcsin [OF assms] arccos [OF assms] apply (auto simp: algebra_simps sin_diff) done then show ?thesis by (simp add: algebra_simps) qed lemma arcsin_arccos_eq: "-1 \ x \ x \ 1 \ arcsin x = pi/2 - arccos x" using arcsin_plus_arccos by force lemma arccos_arcsin_eq: "-1 \ x \ x \ 1 \ arccos x = pi/2 - arcsin x" using arcsin_plus_arccos by force lemma arcsin_arctan: "-1 < x \ x < 1 \ arcsin x = arctan(x / sqrt(1 - x\<^sup>2))" by (simp add: arccos_arctan arcsin_arccos_eq) lemma csqrt_1_diff_eq: "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by ( simp add: of_real_sqrt del: csqrt_of_real_nonneg) lemma arcsin_arccos_sqrt_pos: "0 \ x \ x \ 1 \ arcsin x = arccos(sqrt(1 - x\<^sup>2))" apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arcsin_Arccos_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done lemma arcsin_arccos_sqrt_neg: "-1 \ x \ x \ 0 \ arcsin x = -arccos(sqrt(1 - x\<^sup>2))" using arcsin_arccos_sqrt_pos [of "-x"] by (simp add: arcsin_minus) lemma arccos_arcsin_sqrt_pos: "0 \ x \ x \ 1 \ arccos x = arcsin(sqrt(1 - x\<^sup>2))" apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arccos_Arcsin_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done lemma arccos_arcsin_sqrt_neg: "-1 \ x \ x \ 0 \ arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))" using arccos_arcsin_sqrt_pos [of "-x"] by (simp add: arccos_minus) subsection\<^marker>\tag unimportant\\Continuity results for arcsin and arccos\ lemma continuous_on_Arcsin_real [continuous_intros]: "continuous_on {w \ \. \Re w\ \ 1} Arcsin" proof - have "continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (arcsin (Re x))) = continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (Re (Arcsin (of_real (Re x)))))" by (rule continuous_on_cong [OF refl]) (simp add: arcsin_eq_Re_Arcsin) also have "... = ?thesis" by (rule continuous_on_cong [OF refl]) simp finally show ?thesis using continuous_on_arcsin [OF continuous_on_Re [OF continuous_on_id], of "{w \ \. \Re w\ \ 1}"] continuous_on_of_real by fastforce qed lemma continuous_within_Arcsin_real: "continuous (at z within {w \ \. \Re w\ \ 1}) Arcsin" proof (cases "z \ {w \ \. \Re w\ \ 1}") case True then show ?thesis using continuous_on_Arcsin_real continuous_on_eq_continuous_within by blast next case False with closed_real_abs_le [of 1] show ?thesis by (rule continuous_within_closed_nontrivial) qed lemma continuous_on_Arccos_real: "continuous_on {w \ \. \Re w\ \ 1} Arccos" proof - have "continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (arccos (Re x))) = continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (Re (Arccos (of_real (Re x)))))" by (rule continuous_on_cong [OF refl]) (simp add: arccos_eq_Re_Arccos) also have "... = ?thesis" by (rule continuous_on_cong [OF refl]) simp finally show ?thesis using continuous_on_arccos [OF continuous_on_Re [OF continuous_on_id], of "{w \ \. \Re w\ \ 1}"] continuous_on_of_real by fastforce qed lemma continuous_within_Arccos_real: "continuous (at z within {w \ \. \Re w\ \ 1}) Arccos" proof (cases "z \ {w \ \. \Re w\ \ 1}") case True then show ?thesis using continuous_on_Arccos_real continuous_on_eq_continuous_within by blast next case False with closed_real_abs_le [of 1] show ?thesis by (rule continuous_within_closed_nontrivial) qed lemma sinh_ln_complex: "x \ 0 \ sinh (ln x :: complex) = (x - inverse x) / 2" by (simp add: sinh_def exp_minus scaleR_conv_of_real exp_of_real) lemma cosh_ln_complex: "x \ 0 \ cosh (ln x :: complex) = (x + inverse x) / 2" by (simp add: cosh_def exp_minus scaleR_conv_of_real) lemma tanh_ln_complex: "x \ 0 \ tanh (ln x :: complex) = (x ^ 2 - 1) / (x ^ 2 + 1)" by (simp add: tanh_def sinh_ln_complex cosh_ln_complex divide_simps power2_eq_square) subsection\Roots of unity\ theorem complex_root_unity: fixes j::nat assumes "n \ 0" shows "exp(2 * of_real pi * \ * of_nat j / of_nat n)^n = 1" proof - have *: "of_nat j * (complex_of_real pi * 2) = complex_of_real (2 * real j * pi)" by (simp add: of_real_numeral) then show ?thesis apply (simp add: exp_of_nat_mult [symmetric] mult_ac exp_Euler) apply (simp only: * cos_of_real sin_of_real) apply (simp add: ) done qed lemma complex_root_unity_eq: fixes j::nat and k::nat assumes "1 \ n" shows "(exp(2 * of_real pi * \ * of_nat j / of_nat n) = exp(2 * of_real pi * \ * of_nat k / of_nat n) \ j mod n = k mod n)" proof - have "(\z::int. \ * (of_nat j * (of_real pi * 2)) = \ * (of_nat k * (of_real pi * 2)) + \ * (of_int z * (of_nat n * (of_real pi * 2)))) \ (\z::int. of_nat j * (\ * (of_real pi * 2)) = (of_nat k + of_nat n * of_int z) * (\ * (of_real pi * 2)))" by (simp add: algebra_simps) also have "... \ (\z::int. of_nat j = of_nat k + of_nat n * (of_int z :: complex))" by simp also have "... \ (\z::int. of_nat j = of_nat k + of_nat n * z)" apply (rule HOL.iff_exI) apply (auto simp: ) using of_int_eq_iff apply fastforce by (metis of_int_add of_int_mult of_int_of_nat_eq) also have "... \ int j mod int n = int k mod int n" by (auto simp: mod_eq_dvd_iff dvd_def algebra_simps) also have "... \ j mod n = k mod n" by (metis of_nat_eq_iff zmod_int) finally have "(\z. \ * (of_nat j * (of_real pi * 2)) = \ * (of_nat k * (of_real pi * 2)) + \ * (of_int z * (of_nat n * (of_real pi * 2)))) \ j mod n = k mod n" . note * = this show ?thesis using assms by (simp add: exp_eq field_split_simps mult_ac of_real_numeral *) qed corollary bij_betw_roots_unity: "bij_betw (\j. exp(2 * of_real pi * \ * of_nat j / of_nat n)) {.. * of_nat j / of_nat n) | j. j < n}" by (auto simp: bij_betw_def inj_on_def complex_root_unity_eq) lemma complex_root_unity_eq_1: fixes j::nat and k::nat assumes "1 \ n" shows "exp(2 * of_real pi * \ * of_nat j / of_nat n) = 1 \ n dvd j" proof - have "1 = exp(2 * of_real pi * \ * (of_nat n / of_nat n))" using assms by simp then have "exp(2 * of_real pi * \ * (of_nat j / of_nat n)) = 1 \ j mod n = n mod n" using complex_root_unity_eq [of n j n] assms by simp then show ?thesis by auto qed lemma finite_complex_roots_unity_explicit: "finite {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n}" by simp lemma card_complex_roots_unity_explicit: "card {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n} = n" by (simp add: Finite_Set.bij_betw_same_card [OF bij_betw_roots_unity, symmetric]) lemma complex_roots_unity: assumes "1 \ n" shows "{z::complex. z^n = 1} = {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n}" apply (rule Finite_Set.card_seteq [symmetric]) using assms apply (auto simp: card_complex_roots_unity_explicit finite_roots_unity complex_root_unity card_roots_unity) done lemma card_complex_roots_unity: "1 \ n \ card {z::complex. z^n = 1} = n" by (simp add: card_complex_roots_unity_explicit complex_roots_unity) lemma complex_not_root_unity: "1 \ n \ \u::complex. norm u = 1 \ u^n \ 1" apply (rule_tac x="exp (of_real pi * \ * of_real (1 / n))" in exI) apply (auto simp: Re_complex_div_eq_0 exp_of_nat_mult [symmetric] mult_ac exp_Euler) done + + end diff --git a/src/HOL/Analysis/Conformal_Mappings.thy b/src/HOL/Analysis/Conformal_Mappings.thy --- a/src/HOL/Analysis/Conformal_Mappings.thy +++ b/src/HOL/Analysis/Conformal_Mappings.thy @@ -1,5085 +1,5084 @@ section \Conformal Mappings and Consequences of Cauchy's Integral Theorem\ text\By John Harrison et al. Ported from HOL Light by L C Paulson (2016)\ text\Also Cauchy's residue theorem by Wenda Li (2016)\ theory Conformal_Mappings -imports Cauchy_Integral_Theorem - +imports Cauchy_Integral_Formula begin -(* FIXME mv to Cauchy_Integral_Theorem.thy *) -subsection\Cauchy's inequality and more versions of Liouville\ +subsection\Liouville's theorem\ lemma Cauchy_higher_deriv_bound: assumes holf: "f holomorphic_on (ball z r)" and contf: "continuous_on (cball z r) f" and fin : "\w. w \ ball z r \ f w \ ball y B0" and "0 < r" and "0 < n" shows "norm ((deriv ^^ n) f z) \ (fact n) * B0 / r^n" proof - have "0 < B0" using \0 < r\ fin [of z] by (metis ball_eq_empty ex_in_conv fin not_less) have le_B0: "\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 "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w) z - (deriv ^^ n) (\w. y) z" using \0 < n\ by simp also have "... = (deriv ^^ n) (\w. f w - y) z" by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \0 < r\) finally have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w - y) z" . have contf': "continuous_on (cball z r) (\u. f u - y)" by (rule contf continuous_intros)+ have holf': "(\u. (f u - y)) holomorphic_on (ball z r)" by (simp add: holf holomorphic_on_diff) define a where "a = (2 * pi)/(fact n)" have "0 < a" by (simp add: a_def) have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" using \0 < r\ by (simp add: a_def field_split_simps) have der_dif: "(deriv ^^ n) (\w. f w - y) z = (deriv ^^ n) f z" using \0 < r\ \0 < n\ by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) have "norm ((2 * of_real pi * \)/(fact n) * (deriv ^^ n) (\w. f w - y) z) \ (B0/r^(Suc n)) * (2 * pi * r)" apply (rule has_contour_integral_bound_circlepath [of "(\u. (f u - y)/(u - z)^(Suc n))" _ z]) using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] using \0 < B0\ \0 < r\ apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) done then show ?thesis using \0 < r\ by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) qed + lemma Cauchy_inequality: assumes holf: "f holomorphic_on (ball \ r)" and contf: "continuous_on (cball \ r) f" and "0 < r" and nof: "\x. norm(\-x) = r \ norm(f x) \ B" shows "norm ((deriv ^^ n) f \) \ (fact n) * B / r^n" proof - obtain x where "norm (\-x) = r" by (metis abs_of_nonneg add_diff_cancel_left' \0 < r\ diff_add_cancel dual_order.strict_implies_order norm_of_real) then have "0 \ B" by (metis nof norm_not_less_zero not_le order_trans) have "((\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]) 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 show ?thesis using \0 < r\ by (simp add: norm_divide norm_mult field_simps) qed lemma Liouville_polynomial: assumes holf: "f holomorphic_on UNIV" and nof: "\z. A \ norm z \ norm(f z) \ B * norm z ^ n" shows "f \ = (\k\n. (deriv^^k) f 0 / fact k * \ ^ k)" proof (cases rule: le_less_linear [THEN disjE]) assume "B \ 0" then have "\z. A \ norm z \ norm(f z) = 0" by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le) then have f0: "(f \ 0) at_infinity" using Lim_at_infinity by force then have [simp]: "f = (\w. 0)" using Liouville_weak [OF holf, of 0] by (simp add: eventually_at_infinity f0) meson show ?thesis by simp next assume "0 < B" have "((\k. (deriv ^^ k) f 0 / (fact k) * (\ - 0)^k) sums f \)" apply (rule holomorphic_power_series [where r = "norm \ + 1"]) using holf holomorphic_on_subset apply auto done then have sumsf: "((\k. (deriv ^^ k) f 0 / (fact k) * \^k) sums f \)" by simp have "(deriv ^^ k) f 0 / fact k * \ ^ k = 0" if "k>n" for k proof (cases "(deriv ^^ k) f 0 = 0") case True then show ?thesis by simp next case False define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" have "1 \ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" using \0 < B\ by simp then have wge1: "1 \ norm w" by (metis norm_of_real w_def) then have "w \ 0" by auto have kB: "0 < fact k * B" using \0 < B\ by simp then have "0 \ fact k * B / cmod ((deriv ^^ k) f 0)" by simp then have wgeA: "A \ cmod w" by (simp only: w_def norm_of_real) have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" using \0 < B\ by simp then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" by (metis norm_of_real w_def) then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" using False by (simp add: field_split_simps mult.commute split: if_split_asm) also have "... \ fact k * (B * norm w ^ n) / norm w ^ k" 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) 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 also have "... = fact k * B / cmod w ^ (k-n)" by simp finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . then have "1 / cmod w < 1 / cmod w ^ (k - n)" by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) then have "cmod w ^ (k - n) < cmod w" by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) with self_le_power [OF wge1] have False by (meson diff_is_0_eq not_gr0 not_le that) then show ?thesis by blast qed then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \ ^ (k + Suc n) = 0" for k using not_less_eq by blast then have "(\i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \ ^ (i + Suc n)) sums 0" by (rule sums_0) with sums_split_initial_segment [OF sumsf, where n = "Suc n"] show ?thesis using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce qed text\Every bounded entire function is a constant function.\ theorem Liouville_theorem: assumes holf: "f holomorphic_on UNIV" and bf: "bounded (range f)" obtains c where "\z. f z = c" proof - obtain B where "\z. cmod (f z) \ B" by (meson bf bounded_pos rangeI) then show ?thesis using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast qed text\A holomorphic function f has only isolated zeros unless f is 0.\ lemma powser_0_nonzero: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes r: "0 < r" and sm: "\x. norm (x - \) < r \ (\n. a n * (x - \) ^ n) sums (f x)" and [simp]: "f \ = 0" and m0: "a m \ 0" and "m>0" obtains s where "0 < s" and "\z. z \ cball \ s - {\} \ f z \ 0" proof - have "r \ conv_radius a" using sm sums_summable by (auto simp: le_conv_radius_iff [where \=\]) obtain m where am: "a m \ 0" and az [simp]: "(\n. n a n = 0)" apply (rule_tac m = "LEAST n. a n \ 0" in that) using m0 apply (rule LeastI2) apply (fastforce intro: dest!: not_less_Least)+ done define b where "b i = a (i+m) / a m" for i define g where "g x = suminf (\i. b i * (x - \) ^ i)" for x have [simp]: "b 0 = 1" by (simp add: am b_def) { fix x::'a assume "norm (x - \) < r" then have "(\n. (a m * (x - \)^m) * (b n * (x - \)^n)) sums (f x)" using am az sm sums_zero_iff_shift [of m "(\n. a n * (x - \) ^ n)" "f x"] by (simp add: b_def monoid_mult_class.power_add algebra_simps) then have "x \ \ \ (\n. b n * (x - \)^n) sums (f x / (a m * (x - \)^m))" using am by (simp add: sums_mult_D) } note bsums = this then have "norm (x - \) < r \ summable (\n. b n * (x - \)^n)" for x using sums_summable by (cases "x=\") auto then have "r \ conv_radius b" by (simp add: le_conv_radius_iff [where \=\]) then have "r/2 < conv_radius b" using not_le order_trans r by fastforce then have "continuous_on (cball \ (r/2)) g" using powser_continuous_suminf [of "r/2" b \] by (simp add: g_def) then obtain s where "s>0" "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ dist (g x) (g \) < 1/2" apply (rule continuous_onE [where x=\ and e = "1/2"]) using r apply (auto simp: norm_minus_commute dist_norm) done moreover have "g \ = 1" by (simp add: g_def) ultimately have gnz: "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ (g x) \ 0" by fastforce have "f x \ 0" if "x \ \" "norm (x - \) \ s" "norm (x - \) \ r/2" for x using bsums [of x] that gnz [of x] apply (auto simp: g_def) using r sums_iff by fastforce then show ?thesis apply (rule_tac s="min s (r/2)" in that) using \0 < r\ \0 < s\ by (auto simp: dist_commute dist_norm) qed subsection \Analytic continuation\ proposition isolated_zeros: assumes holf: "f holomorphic_on S" and "open S" "connected S" "\ \ S" "f \ = 0" "\ \ S" "f \ \ 0" obtains r where "0 < r" and "ball \ r \ S" and "\z. z \ ball \ r - {\} \ f z \ 0" proof - obtain r where "0 < r" and r: "ball \ r \ S" using \open S\ \\ \ S\ open_contains_ball_eq by blast have powf: "((\n. (deriv ^^ n) f \ / (fact n) * (z - \)^n) sums f z)" if "z \ ball \ r" for z apply (rule holomorphic_power_series [OF _ that]) apply (rule holomorphic_on_subset [OF holf r]) done obtain m where m: "(deriv ^^ m) f \ / (fact m) \ 0" using holomorphic_fun_eq_0_on_connected [OF holf \open S\ \connected S\ _ \\ \ S\ \\ \ S\] \f \ \ 0\ by auto then have "m \ 0" using assms(5) funpow_0 by fastforce obtain s where "0 < s" and s: "\z. z \ cball \ s - {\} \ f z \ 0" apply (rule powser_0_nonzero [OF \0 < r\ powf \f \ = 0\ m]) using \m \ 0\ by (auto simp: dist_commute dist_norm) have "0 < min r s" by (simp add: \0 < r\ \0 < s\) then show ?thesis apply (rule that) using r s by auto qed proposition analytic_continuation: assumes holf: "f holomorphic_on S" and "open S" and "connected S" and "U \ S" and "\ \ S" and "\ islimpt U" and fU0 [simp]: "\z. z \ U \ f z = 0" and "w \ S" shows "f w = 0" proof - obtain e where "0 < e" and e: "cball \ e \ S" using \open S\ \\ \ S\ open_contains_cball_eq by blast define T where "T = cball \ e \ U" have contf: "continuous_on (closure T) f" by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on holomorphic_on_subset inf.cobounded1) have fT0 [simp]: "\x. x \ T \ f x = 0" by (simp add: T_def) have "\r. \\e>0. \x'\U. x' \ \ \ dist x' \ < e; 0 < r\ \ \x'\cball \ e \ U. x' \ \ \ dist x' \ < r" by (metis \0 < e\ IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj) then have "\ islimpt T" using \\ islimpt U\ by (auto simp: T_def islimpt_approachable) then have "\ \ closure T" by (simp add: closure_def) then have "f \ = 0" by (auto simp: continuous_constant_on_closure [OF contf]) show ?thesis apply (rule ccontr) apply (rule isolated_zeros [OF holf \open S\ \connected S\ \\ \ S\ \f \ = 0\ \w \ S\], assumption) by (metis open_ball \\ islimpt T\ centre_in_ball fT0 insertE insert_Diff islimptE) qed corollary analytic_continuation_open: assumes "open s" and "open s'" and "s \ {}" and "connected s'" and "s \ s'" assumes "f holomorphic_on s'" and "g holomorphic_on s'" and "\z. z \ s \ f z = g z" assumes "z \ s'" shows "f z = g z" proof - from \s \ {}\ obtain \ where "\ \ s" by auto with \open s\ have \: "\ islimpt s" by (intro interior_limit_point) (auto simp: interior_open) have "f z - g z = 0" by (rule analytic_continuation[of "\z. f z - g z" s' s \]) (insert assms \\ \ s\ \, auto intro: holomorphic_intros) thus ?thesis by simp qed subsection\Open mapping theorem\ lemma holomorphic_contract_to_zero: assumes contf: "continuous_on (cball \ r) f" and holf: "f holomorphic_on ball \ r" and "0 < r" and norm_less: "\z. norm(\ - z) = r \ norm(f \) < norm(f z)" obtains z where "z \ ball \ r" "f z = 0" proof - { assume fnz: "\w. w \ ball \ r \ f w \ 0" then have "0 < norm (f \)" by (simp add: \0 < r\) have fnz': "\w. w \ cball \ r \ f w \ 0" by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) have "frontier(cball \ r) \ {}" using \0 < r\ by simp define g where [abs_def]: "g z = inverse (f z)" for z have contg: "continuous_on (cball \ r) g" unfolding g_def using contf continuous_on_inverse fnz' by blast have holg: "g holomorphic_on ball \ r" unfolding g_def using fnz holf holomorphic_on_inverse by blast have "frontier (cball \ r) \ cball \ r" by (simp add: subset_iff) then have contf': "continuous_on (frontier (cball \ r)) f" and contg': "continuous_on (frontier (cball \ r)) g" by (blast intro: contf contg continuous_on_subset)+ have froc: "frontier(cball \ r) \ {}" using \0 < r\ by simp moreover have "continuous_on (frontier (cball \ r)) (norm o f)" using contf' continuous_on_compose continuous_on_norm_id by blast ultimately obtain w where w: "w \ frontier(cball \ r)" and now: "\x. x \ frontier(cball \ r) \ norm (f w) \ norm (f x)" apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]]) apply simp done then have fw: "0 < norm (f w)" by (simp add: fnz') have "continuous_on (frontier (cball \ r)) (norm o g)" using contg' continuous_on_compose continuous_on_norm_id by blast then obtain v where v: "v \ frontier(cball \ r)" and nov: "\x. x \ frontier(cball \ r) \ norm (g v) \ norm (g x)" apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]]) apply simp done then have fv: "0 < norm (f v)" by (simp add: fnz') have "norm ((deriv ^^ 0) g \) \ fact 0 * norm (g v) / r ^ 0" by (rule Cauchy_inequality [OF holg contg \0 < r\]) (simp add: dist_norm nov) then have "cmod (g \) \ norm (g v)" by simp with w have wr: "norm (\ - w) = r" and nfw: "norm (f w) \ norm (f \)" apply (simp_all add: dist_norm) by (metis \0 < cmod (f \)\ g_def less_imp_inverse_less norm_inverse not_le now order_trans v) with fw have False using norm_less by force } with that show ?thesis by blast qed theorem open_mapping_thm: assumes holf: "f holomorphic_on S" and S: "open S" and "connected S" and "open U" and "U \ S" and fne: "\ f constant_on S" shows "open (f ` U)" proof - have *: "open (f ` U)" if "U \ {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\x. \y \ U. f y \ x" for U proof (clarsimp simp: open_contains_ball) fix \ assume \: "\ \ U" show "\e>0. ball (f \) e \ f ` U" proof - have hol: "(\z. f z - f \) holomorphic_on U" by (rule holomorphic_intros that)+ obtain s where "0 < s" and sbU: "ball \ s \ U" and sne: "\z. z \ ball \ s - {\} \ (\z. f z - f \) z \ 0" using isolated_zeros [OF hol U \] by (metis fneU right_minus_eq) obtain r where "0 < r" and r: "cball \ r \ ball \ s" apply (rule_tac r="s/2" in that) using \0 < s\ by auto have "cball \ r \ U" using sbU r by blast then have frsbU: "frontier (cball \ r) \ U" using Diff_subset frontier_def order_trans by fastforce then have cof: "compact (frontier(cball \ r))" by blast have frne: "frontier (cball \ r) \ {}" using \0 < r\ by auto have contfr: "continuous_on (frontier (cball \ r)) (\z. norm (f z - f \))" by (metis continuous_on_norm continuous_on_subset frsbU hol holomorphic_on_imp_continuous_on) obtain w where "norm (\ - w) = r" and w: "(\z. norm (\ - z) = r \ norm (f w - f \) \ norm(f z - f \))" apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]]) apply (simp add: dist_norm) done moreover define \ where "\ \ norm (f w - f \) / 3" ultimately have "0 < \" using \0 < r\ dist_complex_def r sne by auto have "ball (f \) \ \ f ` U" proof fix \ assume \: "\ \ ball (f \) \" have *: "cmod (\ - f \) < cmod (\ - f z)" if "cmod (\ - z) = r" for z proof - have lt: "cmod (f w - f \) / 3 < cmod (\ - f z)" using w [OF that] \ using dist_triangle2 [of "f \" "\" "f z"] dist_triangle2 [of "f \" "f z" \] by (simp add: \_def dist_norm norm_minus_commute) show ?thesis by (metis \_def dist_commute dist_norm less_trans lt mem_ball \) qed have "continuous_on (cball \ r) (\z. \ - f z)" apply (rule continuous_intros)+ using \cball \ r \ U\ \f holomorphic_on U\ apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) done moreover have "(\z. \ - f z) holomorphic_on ball \ r" apply (rule holomorphic_intros)+ apply (metis \cball \ r \ U\ \f holomorphic_on U\ holomorphic_on_subset interior_cball interior_subset) done ultimately obtain z where "z \ ball \ r" "\ - f z = 0" apply (rule holomorphic_contract_to_zero) apply (blast intro!: \0 < r\ *)+ done then show "\ \ f ` U" using \cball \ r \ U\ by fastforce qed then show ?thesis using \0 < \\ by blast qed qed have "open (f ` X)" if "X \ components U" for X proof - have holfU: "f holomorphic_on U" using \U \ S\ holf holomorphic_on_subset by blast have "X \ {}" using that by (simp add: in_components_nonempty) moreover have "open X" using that \open U\ open_components by auto moreover have "connected X" using that in_components_maximal by blast moreover have "f holomorphic_on X" by (meson that holfU holomorphic_on_subset in_components_maximal) moreover have "\y\X. f y \ x" for x proof (rule ccontr) assume not: "\ (\y\X. f y \ x)" have "X \ S" using \U \ S\ in_components_subset that by blast obtain w where w: "w \ X" using \X \ {}\ by blast have wis: "w islimpt X" using w \open X\ interior_eq by auto have hol: "(\z. f z - x) holomorphic_on S" by (simp add: holf holomorphic_on_diff) with fne [unfolded constant_on_def] analytic_continuation[OF hol S \connected S\ \X \ S\ _ wis] not \X \ S\ w show False by auto qed ultimately show ?thesis by (rule *) qed then have "open (f ` \(components U))" by (metis (no_types, lifting) imageE image_Union open_Union) then show ?thesis by force qed text\No need for \<^term>\S\ to be connected. But the nonconstant condition is stronger.\ corollary\<^marker>\tag unimportant\ open_mapping_thm2: assumes holf: "f holomorphic_on S" and S: "open S" and "open U" "U \ S" and fnc: "\X. \open X; X \ S; X \ {}\ \ \ f constant_on X" shows "open (f ` U)" proof - have "S = \(components S)" by simp with \U \ S\ have "U = (\C \ components S. C \ U)" by auto then have "f ` U = (\C \ components S. f ` (C \ U))" using image_UN by fastforce moreover { fix C assume "C \ components S" with S \C \ components S\ open_components in_components_connected have C: "open C" "connected C" by auto have "C \ S" by (metis \C \ components S\ in_components_maximal) have nf: "\ f constant_on C" apply (rule fnc) using C \C \ S\ \C \ components S\ in_components_nonempty by auto have "f holomorphic_on C" by (metis holf holomorphic_on_subset \C \ S\) then have "open (f ` (C \ U))" apply (rule open_mapping_thm [OF _ C _ _ nf]) apply (simp add: C \open U\ open_Int, blast) done } ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ open_mapping_thm3: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" shows "open (f ` S)" apply (rule open_mapping_thm2 [OF holf]) using assms apply (simp_all add:) using injective_not_constant subset_inj_on by blast subsection\Maximum modulus principle\ text\If \<^term>\f\ is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is properly within the domain of \<^term>\f\.\ proposition maximum_modulus_principle: assumes holf: "f holomorphic_on S" and S: "open S" and "connected S" and "open U" and "U \ S" and "\ \ U" and no: "\z. z \ U \ norm(f z) \ norm(f \)" shows "f constant_on S" proof (rule ccontr) assume "\ f constant_on S" then have "open (f ` U)" using open_mapping_thm assms by blast moreover have "\ open (f ` U)" proof - have "\t. cmod (f \ - t) < e \ t \ f ` U" if "0 < e" for e apply (rule_tac x="if 0 < Re(f \) then f \ + (e/2) else f \ - (e/2)" in exI) using that apply (simp add: dist_norm) apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym) done then show ?thesis unfolding open_contains_ball by (metis \\ \ U\ contra_subsetD dist_norm imageI mem_ball) qed ultimately show False by blast qed proposition maximum_modulus_frontier: assumes holf: "f holomorphic_on (interior S)" and contf: "continuous_on (closure S) f" and bos: "bounded S" and leB: "\z. z \ frontier S \ norm(f z) \ B" and "\ \ S" shows "norm(f \) \ B" proof - have "compact (closure S)" using bos by (simp add: bounded_closure compact_eq_bounded_closed) moreover have "continuous_on (closure S) (cmod \ f)" using contf continuous_on_compose continuous_on_norm_id by blast ultimately obtain z where zin: "z \ closure S" and z: "\y. y \ closure S \ (cmod \ f) y \ (cmod \ f) z" using continuous_attains_sup [of "closure S" "norm o f"] \\ \ S\ by auto then consider "z \ frontier S" | "z \ interior S" using frontier_def by auto then have "norm(f z) \ B" proof cases case 1 then show ?thesis using leB by blast next case 2 have zin: "z \ connected_component_set (interior S) z" by (simp add: 2) have "f constant_on (connected_component_set (interior S) z)" apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin]) apply (metis connected_component_subset holf holomorphic_on_subset) apply (simp_all add: open_connected_component) by (metis closure_subset comp_eq_dest_lhs interior_subset subsetCE z connected_component_in) then obtain c where c: "\w. w \ connected_component_set (interior S) z \ f w = c" by (auto simp: constant_on_def) have "f ` closure(connected_component_set (interior S) z) \ {c}" apply (rule image_closure_subset) apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset) using c apply auto done then have cc: "\w. w \ closure(connected_component_set (interior S) z) \ f w = c" by blast have "frontier(connected_component_set (interior S) z) \ {}" apply (simp add: frontier_eq_empty) by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV) then obtain w where w: "w \ frontier(connected_component_set (interior S) z)" by auto then have "norm (f z) = norm (f w)" by (simp add: "2" c cc frontier_def) also have "... \ B" apply (rule leB) using w using frontier_interior_subset frontier_of_connected_component_subset by blast finally show ?thesis . qed then show ?thesis using z \\ \ S\ closure_subset by fastforce qed corollary\<^marker>\tag unimportant\ maximum_real_frontier: assumes holf: "f holomorphic_on (interior S)" and contf: "continuous_on (closure S) f" and bos: "bounded S" and leB: "\z. z \ frontier S \ Re(f z) \ B" and "\ \ S" shows "Re(f \) \ B" using maximum_modulus_frontier [of "exp o f" S "exp B"] Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms by auto subsection\<^marker>\tag unimportant\ \Factoring out a zero according to its order\ lemma holomorphic_factor_order_of_zero: assumes holf: "f holomorphic_on S" and os: "open S" and "\ \ S" "0 < n" and dnz: "(deriv ^^ n) f \ \ 0" and dfz: "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" obtains g r where "0 < r" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" "\w. w \ ball \ r \ g w \ 0" proof - obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) then have holfb: "f holomorphic_on ball \ r" using holf holomorphic_on_subset by blast define g where "g w = suminf (\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i)" for w have sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" and feq: "f w - f \ = (w - \)^n * g w" if w: "w \ ball \ r" for w proof - define powf where "powf = (\i. (deriv ^^ i) f \/(fact i) * (w - \)^i)" have sing: "{.. = 0 then {} else {0})" unfolding powf_def using \0 < n\ dfz by (auto simp: dfz; metis funpow_0 not_gr0) have "powf sums f w" unfolding powf_def by (rule holomorphic_power_series [OF holfb w]) moreover have "(\i" apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric]) apply simp apply (simp only: dfz sing) apply (simp add: powf_def) done ultimately have fsums: "(\i. powf (i+n)) sums (f w - f \)" using w sums_iff_shift' by metis then have *: "summable (\i. (w - \) ^ n * ((deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n)))" unfolding powf_def using sums_summable by (auto simp: power_add mult_ac) have "summable (\i. (deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n))" proof (cases "w=\") case False then show ?thesis using summable_mult [OF *, of "1 / (w - \) ^ n"] by simp next case True then show ?thesis by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"] split: if_split_asm) qed then show sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" by (simp add: summable_sums_iff g_def) show "f w - f \ = (w - \)^n * g w" apply (rule sums_unique2) apply (rule fsums [unfolded powf_def]) using sums_mult [OF sumsg, of "(w - \) ^ n"] by (auto simp: power_add mult_ac) qed then have holg: "g holomorphic_on ball \ r" by (meson sumsg power_series_holomorphic) then have contg: "continuous_on (ball \ r) g" by (blast intro: holomorphic_on_imp_continuous_on) have "g \ \ 0" using dnz unfolding g_def by (subst suminf_finite [of "{0}"]) auto obtain d where "0 < d" and d: "\w. w \ ball \ d \ g w \ 0" apply (rule exE [OF continuous_on_avoid [OF contg _ \g \ \ 0\]]) using \0 < r\ apply force by (metis \0 < r\ less_trans mem_ball not_less_iff_gr_or_eq) show ?thesis apply (rule that [where g=g and r ="min r d"]) using \0 < r\ \0 < d\ holg apply (auto simp: feq holomorphic_on_subset subset_ball d) done qed lemma holomorphic_factor_order_of_zero_strong: assumes holf: "f holomorphic_on S" "open S" "\ \ S" "0 < n" and "(deriv ^^ n) f \ \ 0" and "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" obtains g r where "0 < r" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w - f \ = ((w - \) * g w) ^ n" "\w. w \ ball \ r \ g w \ 0" proof - obtain g r where "0 < r" and holg: "g holomorphic_on ball \ r" and feq: "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" and gne: "\w. w \ ball \ r \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF assms]) have con: "continuous_on (ball \ r) (\z. deriv g z / g z)" by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on) have cd: "\x. dist \ x < r \ (\z. deriv g z / g z) field_differentiable at x" apply (rule derivative_intros)+ using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) apply (metis open_ball at_within_open holg holomorphic_on_def mem_ball) using gne mem_ball by blast obtain h where h: "\x. x \ ball \ r \ (h has_field_derivative deriv g x / g x) (at x)" apply (rule exE [OF holomorphic_convex_primitive [of "ball \ r" "{}" "\z. deriv g z / g z"]]) apply (auto simp: con cd) apply (metis open_ball at_within_open mem_ball) done then have "continuous_on (ball \ r) h" by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open) then have con: "continuous_on (ball \ r) (\x. exp (h x) / g x)" by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne) have 0: "dist \ x < r \ ((\x. exp (h x) / g x) has_field_derivative 0) (at x)" for x apply (rule h derivative_eq_intros | simp)+ apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2]) using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h) done obtain c where c: "\x. x \ ball \ r \ exp (h x) / g x = c" by (rule DERIV_zero_connected_constant [of "ball \ r" "{}" "\x. exp(h x) / g x"]) (auto simp: con 0) have hol: "(\z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \ r" apply (rule holomorphic_on_compose [unfolded o_def, where g = exp]) apply (rule holomorphic_intros)+ using h holomorphic_on_open apply blast apply (rule holomorphic_intros)+ using \0 < n\ apply simp apply (rule holomorphic_intros)+ done show ?thesis apply (rule that [where g="\z. exp((Ln(inverse c) + h z)/n)" and r =r]) using \0 < r\ \0 < n\ apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric]) apply (rule hol) apply (simp add: Transcendental.exp_add gne) done qed lemma fixes k :: "'a::wellorder" assumes a_def: "a == LEAST x. P x" and P: "P k" shows def_LeastI: "P a" and def_Least_le: "a \ k" unfolding a_def by (rule LeastI Least_le; rule P)+ lemma holomorphic_factor_zero_nonconstant: assumes holf: "f holomorphic_on S" and S: "open S" "connected S" and "\ \ S" "f \ = 0" and nonconst: "\ f constant_on S" obtains g r n where "0 < n" "0 < r" "ball \ r \ S" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w = (w - \)^n * g w" "\w. w \ ball \ r \ g w \ 0" proof (cases "\n>0. (deriv ^^ n) f \ = 0") case True then show ?thesis using holomorphic_fun_eq_const_on_connected [OF holf S _ \\ \ S\] nonconst by (simp add: constant_on_def) next case False then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \ \ 0" by blast obtain r0 where "r0 > 0" "ball \ r0 \ S" using S openE \\ \ S\ by auto define n where "n \ LEAST n. (deriv ^^ n) f \ \ 0" have n_ne: "(deriv ^^ n) f \ \ 0" by (rule def_LeastI [OF n_def]) (rule n0) then have "0 < n" using \f \ = 0\ using funpow_0 by fastforce have n_min: "\k. k < n \ (deriv ^^ k) f \ = 0" using def_Least_le [OF n_def] not_le by blast then obtain g r1 where "0 < r1" "g holomorphic_on ball \ r1" "\w. w \ ball \ r1 \ f w = (w - \) ^ n * g w" "\w. w \ ball \ r1 \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne] simp: \f \ = 0\) then show ?thesis apply (rule_tac g=g and r="min r0 r1" and n=n in that) using \0 < n\ \0 < r0\ \0 < r1\ \ball \ r0 \ S\ apply (auto simp: subset_ball intro: holomorphic_on_subset) done qed lemma holomorphic_lower_bound_difference: assumes holf: "f holomorphic_on S" and S: "open S" "connected S" and "\ \ S" and "\ \ S" and fne: "f \ \ f \" obtains k n r where "0 < k" "0 < r" "ball \ r \ S" "\w. w \ ball \ r \ k * norm(w - \)^n \ norm(f w - f \)" proof - define n where "n = (LEAST n. 0 < n \ (deriv ^^ n) f \ \ 0)" obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \ \ 0" using fne holomorphic_fun_eq_const_on_connected [OF holf S] \\ \ S\ \\ \ S\ by blast then have "0 < n" and n_ne: "(deriv ^^ n) f \ \ 0" unfolding n_def by (metis (mono_tags, lifting) LeastI)+ have n_min: "\k. \0 < k; k < n\ \ (deriv ^^ k) f \ = 0" unfolding n_def by (blast dest: not_less_Least) then obtain g r where "0 < r" and holg: "g holomorphic_on ball \ r" and fne: "\w. w \ ball \ r \ f w - f \ = (w - \) ^ n * g w" and gnz: "\w. w \ ball \ r \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne]) obtain e where "e>0" and e: "ball \ e \ S" using assms by (blast elim!: openE) then have holfb: "f holomorphic_on ball \ e" using holf holomorphic_on_subset by blast define d where "d = (min e r) / 2" have "0 < d" using \0 < r\ \0 < e\ by (simp add: d_def) have "d < r" using \0 < r\ by (auto simp: d_def) then have cbb: "cball \ d \ ball \ r" by (auto simp: cball_subset_ball_iff) then have "g holomorphic_on cball \ d" by (rule holomorphic_on_subset [OF holg]) then have "closed (g ` cball \ d)" by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on) moreover have "g ` cball \ d \ {}" using \0 < d\ by auto ultimately obtain x where x: "x \ g ` cball \ d" and "\y. y \ g ` cball \ d \ dist 0 x \ dist 0 y" by (rule distance_attains_inf) blast then have leg: "\w. w \ cball \ d \ norm x \ norm (g w)" by auto have "ball \ d \ cball \ d" by auto also have "... \ ball \ e" using \0 < d\ d_def by auto also have "... \ S" by (rule e) finally have dS: "ball \ d \ S" . moreover have "x \ 0" using gnz x \d < r\ by auto ultimately show ?thesis apply (rule_tac k="norm x" and n=n and r=d in that) using \d < r\ leg apply (auto simp: \0 < d\ fne norm_mult norm_power algebra_simps mult_right_mono) done qed lemma assumes holf: "f holomorphic_on (S - {\})" and \: "\ \ interior S" shows holomorphic_on_extend_lim: "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ ((\z. (z - \) * f z) \ 0) (at \)" (is "?P = ?Q") and holomorphic_on_extend_bounded: "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ (\B. eventually (\z. norm(f z) \ B) (at \))" (is "?P = ?R") proof - obtain \ where "0 < \" and \: "ball \ \ \ S" using \ mem_interior by blast have "?R" if holg: "g holomorphic_on S" and gf: "\z. z \ S - {\} \ g z = f z" for g proof - have *: "\\<^sub>F z in at \. dist (g z) (g \) < 1 \ cmod (f z) \ cmod (g \) + 1" apply (simp add: eventually_at) apply (rule_tac x="\" in exI) using \ \0 < \\ apply (clarsimp simp:) apply (drule_tac c=x in subsetD) apply (simp add: dist_commute) by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD) have "continuous_on (interior S) g" by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset) then have "\x. x \ interior S \ (g \ g x) (at x)" using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast then have "(g \ g \) (at \)" by (simp add: \) then show ?thesis apply (rule_tac x="norm(g \) + 1" in exI) apply (rule eventually_mp [OF * tendstoD [where e=1]], auto) done qed moreover have "?Q" if "\\<^sub>F z in at \. cmod (f z) \ B" for B by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero) moreover have "?P" if "(\z. (z - \) * f z) \\\ 0" proof - define h where [abs_def]: "h z = (z - \)^2 * f z" for z have h0: "(h has_field_derivative 0) (at \)" apply (simp add: h_def has_field_derivative_iff) apply (rule Lim_transform_within [OF that, of 1]) apply (auto simp: field_split_simps power2_eq_square) done have holh: "h holomorphic_on S" proof (simp add: holomorphic_on_def, clarify) fix z assume "z \ S" show "h field_differentiable at z within S" proof (cases "z = \") case True then show ?thesis using field_differentiable_at_within field_differentiable_def h0 by blast next case False then have "f field_differentiable at z within S" using holomorphic_onD [OF holf, of z] \z \ S\ unfolding field_differentiable_def has_field_derivative_iff by (force intro: exI [where x="dist \ z"] elim: Lim_transform_within_set [unfolded eventually_at]) then show ?thesis by (simp add: h_def power2_eq_square derivative_intros) qed qed define g where [abs_def]: "g z = (if z = \ then deriv h \ else (h z - h \) / (z - \))" for z have holg: "g holomorphic_on S" unfolding g_def by (rule pole_lemma [OF holh \]) show ?thesis apply (rule_tac x="\z. if z = \ then deriv g \ else (g z - g \)/(z - \)" in exI) apply (rule conjI) apply (rule pole_lemma [OF holg \]) apply (auto simp: g_def power2_eq_square divide_simps) using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square) done qed ultimately show "?P = ?Q" and "?P = ?R" by meson+ qed lemma pole_at_infinity: assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \ l) at_infinity" obtains a n where "\z. f z = (\i\n. a i * z^i)" proof (cases "l = 0") case False with tendsto_inverse [OF lim] show ?thesis apply (rule_tac a="(\n. inverse l)" and n=0 in that) apply (simp add: Liouville_weak [OF holf, of "inverse l"]) done next case True then have [simp]: "l = 0" . show ?thesis proof (cases "\r. 0 < r \ (\z \ ball 0 r - {0}. f(inverse z) \ 0)") case True then obtain r where "0 < r" and r: "\z. z \ ball 0 r - {0} \ f(inverse z) \ 0" by auto have 1: "inverse \ f \ inverse holomorphic_on ball 0 r - {0}" by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+ have 2: "0 \ interior (ball 0 r)" using \0 < r\ by simp have "\B. 0 eventually (\z. cmod ((inverse \ f \ inverse) z) \ B) (at 0)" apply (rule exI [where x=1]) apply simp using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] apply (rule eventually_mono) apply (simp add: dist_norm) done with holomorphic_on_extend_bounded [OF 1 2] obtain g where holg: "g holomorphic_on ball 0 r" and geq: "\z. z \ ball 0 r - {0} \ g z = (inverse \ f \ inverse) z" by meson have ifi0: "(inverse \ f \ inverse) \0\ 0" using \l = 0\ lim lim_at_infinity_0 by blast have g2g0: "g \0\ g 0" using \0 < r\ centre_in_ball continuous_at continuous_on_eq_continuous_at holg by (blast intro: holomorphic_on_imp_continuous_on) have g2g1: "g \0\ 0" apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]]) using \0 < r\ by (auto simp: geq) have [simp]: "g 0 = 0" by (rule tendsto_unique [OF _ g2g0 g2g1]) simp have "ball 0 r - {0::complex} \ {}" using \0 < r\ apply (clarsimp simp: ball_def dist_norm) apply (drule_tac c="of_real r/2" in subsetD, auto) done then obtain w::complex where "w \ 0" and w: "norm w < r" by force then have "g w \ 0" by (simp add: geq r) obtain B n e where "0 < B" "0 < e" "e \ r" and leg: "\w. norm w < e \ B * cmod w ^ n \ cmod (g w)" apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w]) using \0 < r\ w \g w \ 0\ by (auto simp: ball_subset_ball_iff) have "cmod (f z) \ cmod z ^ n / B" if "2/e \ cmod z" for z proof - have ize: "inverse z \ ball 0 e - {0}" using that \0 < e\ by (auto simp: norm_divide field_split_simps algebra_simps) then have [simp]: "z \ 0" and izr: "inverse z \ ball 0 r - {0}" using \e \ r\ by auto then have [simp]: "f z \ 0" using r [of "inverse z"] by simp have [simp]: "f z = inverse (g (inverse z))" using izr geq [of "inverse z"] by simp show ?thesis using ize leg [of "inverse z"] \0 < B\ \0 < e\ by (simp add: field_split_simps norm_divide algebra_simps) qed then show ?thesis apply (rule_tac a = "\k. (deriv ^^ k) f 0 / (fact k)" and n=n in that) apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp) done next case False then have fi0: "\r. r > 0 \ \z\ball 0 r - {0}. f (inverse z) = 0" by simp have fz0: "f z = 0" if "0 < r" and lt1: "\x. x \ 0 \ cmod x < r \ inverse (cmod (f (inverse x))) < 1" for z r proof - have f0: "(f \ 0) at_infinity" proof - have DIM_complex[intro]: "2 \ DIM(complex)" \ \should not be necessary!\ by simp from lt1 have "f (inverse x) \ 0 \ x \ 0 \ cmod x < r \ 1 < cmod (f (inverse x))" for x using one_less_inverse by force then have **: "cmod (f (inverse x)) \ 1 \ x \ 0 \ cmod x < r \ f (inverse x) = 0" for x by force then have *: "(f \ inverse) ` (ball 0 r - {0}) \ {0} \ - ball 0 1" by force have "continuous_on (inverse ` (ball 0 r - {0})) f" using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast then have "connected ((f \ inverse) ` (ball 0 r - {0}))" apply (intro connected_continuous_image continuous_intros) apply (force intro: connected_punctured_ball)+ done then have "{0} \ (f \ inverse) ` (ball 0 r - {0}) = {} \ - ball 0 1 \ (f \ inverse) ` (ball 0 r - {0}) = {}" by (rule connected_closedD) (use * in auto) then have "w \ 0 \ cmod w < r \ f (inverse w) = 0" for w using fi0 **[of w] \0 < r\ apply (auto simp add: inf.commute [of "- ball 0 1"] Diff_eq [symmetric] image_subset_iff dest: less_imp_le) apply fastforce apply (drule bspec [of _ _ w]) apply (auto dest: less_imp_le) done then show ?thesis apply (simp add: lim_at_infinity_0) apply (rule tendsto_eventually) apply (simp add: eventually_at) apply (rule_tac x=r in exI) apply (simp add: \0 < r\ dist_norm) done qed obtain w where "w \ ball 0 r - {0}" and "f (inverse w) = 0" using False \0 < r\ by blast then show ?thesis by (auto simp: f0 Liouville_weak [OF holf, of 0]) qed show ?thesis apply (rule that [of "\n. 0" 0]) using lim [unfolded lim_at_infinity_0] apply (simp add: Lim_at dist_norm norm_inverse) apply (drule_tac x=1 in spec) using fz0 apply auto done qed qed subsection\<^marker>\tag unimportant\ \Entire proper functions are precisely the non-trivial polynomials\ lemma proper_map_polyfun: fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" assumes "closed S" and "compact K" and c: "c i \ 0" "1 \ i" "i \ n" shows "compact (S \ {z. (\i\n. c i * z^i) \ K})" proof - obtain B where "B > 0" and B: "\x. x \ K \ norm x \ B" by (metis compact_imp_bounded \compact K\ bounded_pos) have *: "norm x \ b" if "\x. b \ norm x \ B + 1 \ norm (\i\n. c i * x ^ i)" "(\i\n. c i * x ^ i) \ K" for b x proof - have "norm (\i\n. c i * x ^ i) \ B" using B that by blast moreover have "\ B + 1 \ B" by simp ultimately show "norm x \ b" using that by (metis (no_types) less_eq_real_def not_less order_trans) qed have "bounded {z. (\i\n. c i * z ^ i) \ K}" using polyfun_extremal [where c=c and B="B+1", OF c] by (auto simp: bounded_pos eventually_at_infinity_pos *) moreover have "closed ((\z. (\i\n. c i * z ^ i)) -` K)" apply (intro allI continuous_closed_vimage continuous_intros) using \compact K\ compact_eq_bounded_closed by blast ultimately show ?thesis using closed_Int_compact [OF \closed S\] compact_eq_bounded_closed by (auto simp add: vimage_def) qed lemma proper_map_polyfun_univ: fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" assumes "compact K" "c i \ 0" "1 \ i" "i \ n" shows "compact ({z. (\i\n. c i * z^i) \ K})" using proper_map_polyfun [of UNIV K c i n] assms by simp lemma proper_map_polyfun_eq: assumes "f holomorphic_on UNIV" shows "(\k. compact k \ compact {z. f z \ k}) \ (\c n. 0 < n \ (c n \ 0) \ f = (\z. \i\n. c i * z^i))" (is "?lhs = ?rhs") proof assume compf [rule_format]: ?lhs have 2: "\k. 0 < k \ a k \ 0 \ f = (\z. \i \ k. a i * z ^ i)" if "\z. f z = (\i\n. a i * z ^ i)" for a n proof (cases "\i\n. 0 a i = 0") case True then have [simp]: "\z. f z = a 0" by (simp add: that sum.atMost_shift) have False using compf [of "{a 0}"] by simp then show ?thesis .. next case False then obtain k where k: "0 < k" "k\n" "a k \ 0" by force define m where "m = (GREATEST k. k\n \ a k \ 0)" have m: "m\n \ a m \ 0" unfolding m_def apply (rule GreatestI_nat [where b = n]) using k apply auto done have [simp]: "a i = 0" if "m < i" "i \ n" for i using Greatest_le_nat [where b = "n" and P = "\k. k\n \ a k \ 0"] using m_def not_le that by auto have "k \ m" unfolding m_def apply (rule Greatest_le_nat [where b = "n"]) using k apply auto done with k m show ?thesis by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right) qed have "((inverse \ f) \ 0) at_infinity" proof (rule Lim_at_infinityI) fix e::real assume "0 < e" with compf [of "cball 0 (inverse e)"] show "\B. \x. B \ cmod x \ dist ((inverse \ f) x) 0 \ e" apply simp apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse) apply (rule_tac x="b+1" in exI) apply (metis inverse_inverse_eq less_add_same_cancel2 less_imp_inverse_less add.commute not_le not_less_iff_gr_or_eq order_trans zero_less_one) done qed then show ?rhs apply (rule pole_at_infinity [OF assms]) using 2 apply blast done next assume ?rhs then obtain c n where "0 < n" "c n \ 0" "f = (\z. \i\n. c i * z ^ i)" by blast then have "compact {z. f z \ k}" if "compact k" for k by (auto intro: proper_map_polyfun_univ [OF that]) then show ?lhs by blast qed subsection \Relating invertibility and nonvanishing of derivative\ lemma has_complex_derivative_locally_injective: assumes holf: "f holomorphic_on S" and S: "\ \ S" "open S" and dnz: "deriv f \ \ 0" obtains r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" proof - have *: "\d>0. \x. dist \ x < d \ onorm (\v. deriv f x * v - deriv f \ * v) < e" if "e > 0" for e proof - have contdf: "continuous_on S (deriv f)" by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open S\) obtain \ where "\>0" and \: "\x. \x \ S; dist x \ \ \\ \ cmod (deriv f x - deriv f \) \ e/2" using continuous_onE [OF contdf \\ \ S\, of "e/2"] \0 < e\ by (metis dist_complex_def half_gt_zero less_imp_le) obtain \ where "\>0" "ball \ \ \ S" by (metis openE [OF \open S\ \\ \ S\]) with \\>0\ have "\\>0. \x. dist \ x < \ \ onorm (\v. deriv f x * v - deriv f \ * v) \ e/2" apply (rule_tac x="min \ \" in exI) apply (intro conjI allI impI Operator_Norm.onorm_le) apply simp apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult) apply (rule mult_right_mono [OF \]) apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \) done with \e>0\ show ?thesis by force qed have "inj ((*) (deriv f \))" using dnz by simp then obtain g' where g': "linear g'" "g' \ (*) (deriv f \) = id" using linear_injective_left_inverse [of "(*) (deriv f \)"] by (auto simp: linear_times) show ?thesis apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\z h. deriv f z * h" and g' = g']) using g' * apply (simp_all add: linear_conv_bounded_linear that) using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf holomorphic_on_imp_differentiable_at \open S\ apply blast done qed lemma has_complex_derivative_locally_invertible: assumes holf: "f holomorphic_on S" and S: "\ \ S" "open S" and dnz: "deriv f \ \ 0" obtains r where "r > 0" "ball \ r \ S" "open (f ` (ball \ r))" "inj_on f (ball \ r)" proof - obtain r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" by (blast intro: that has_complex_derivative_locally_injective [OF assms]) then have \: "\ \ ball \ r" by simp then have nc: "\ f constant_on ball \ r" using \inj_on f (ball \ r)\ injective_not_constant by fastforce have holf': "f holomorphic_on ball \ r" using \ball \ r \ S\ holf holomorphic_on_subset by blast have "open (f ` ball \ r)" apply (rule open_mapping_thm [OF holf']) using nc apply auto done then show ?thesis using \0 < r\ \ball \ r \ S\ \inj_on f (ball \ r)\ that by blast qed lemma holomorphic_injective_imp_regular: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" and "\ \ S" shows "deriv f \ \ 0" proof - obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) have holf': "f holomorphic_on ball \ r" using \ball \ r \ S\ holf holomorphic_on_subset by blast show ?thesis proof (cases "\n>0. (deriv ^^ n) f \ = 0") case True have fcon: "f w = f \" if "w \ ball \ r" for w apply (rule holomorphic_fun_eq_const_on_connected [OF holf']) using True \0 < r\ that by auto have False using fcon [of "\ + r/2"] \0 < r\ r injf unfolding inj_on_def by (metis \\ \ S\ contra_subsetD dist_commute fcon mem_ball perfect_choose_dist) then show ?thesis .. next case False then obtain n0 where n0: "n0 > 0 \ (deriv ^^ n0) f \ \ 0" by blast define n where [abs_def]: "n = (LEAST n. n > 0 \ (deriv ^^ n) f \ \ 0)" have n_ne: "n > 0" "(deriv ^^ n) f \ \ 0" using def_LeastI [OF n_def n0] by auto have n_min: "\k. 0 < k \ k < n \ (deriv ^^ k) f \ = 0" using def_Least_le [OF n_def] not_le by auto obtain g \ where "0 < \" and holg: "g holomorphic_on ball \ \" and fd: "\w. w \ ball \ \ \ f w - f \ = ((w - \) * g w) ^ n" and gnz: "\w. w \ ball \ \ \ g w \ 0" apply (rule holomorphic_factor_order_of_zero_strong [OF holf \open S\ \\ \ S\ n_ne]) apply (blast intro: n_min)+ done show ?thesis proof (cases "n=1") case True with n_ne show ?thesis by auto next case False have holgw: "(\w. (w - \) * g w) holomorphic_on ball \ (min r \)" apply (rule holomorphic_intros)+ using holg by (simp add: holomorphic_on_subset subset_ball) have gd: "\w. dist \ w < \ \ (g has_field_derivative deriv g w) (at w)" using holg by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH) have *: "\w. w \ ball \ (min r \) \ ((\w. (w - \) * g w) has_field_derivative ((w - \) * deriv g w + g w)) (at w)" by (rule gd derivative_eq_intros | simp)+ have [simp]: "deriv (\w. (w - \) * g w) \ \ 0" using * [of \] \0 < \\ \0 < r\ by (simp add: DERIV_imp_deriv gnz) obtain T where "\ \ T" "open T" and Tsb: "T \ ball \ (min r \)" and oimT: "open ((\w. (w - \) * g w) ` T)" apply (rule has_complex_derivative_locally_invertible [OF holgw, of \]) using \0 < r\ \0 < \\ apply (simp_all add:) by (meson open_ball centre_in_ball) define U where "U = (\w. (w - \) * g w) ` T" have "open U" by (metis oimT U_def) have "0 \ U" apply (auto simp: U_def) apply (rule image_eqI [where x = \]) apply (auto simp: \\ \ T\) done then obtain \ where "\>0" and \: "cball 0 \ \ U" using \open U\ open_contains_cball by blast then have "\ * exp(2 * of_real pi * \ * (0/n)) \ cball 0 \" "\ * exp(2 * of_real pi * \ * (1/n)) \ cball 0 \" by (auto simp: norm_mult) with \ have "\ * exp(2 * of_real pi * \ * (0/n)) \ U" "\ * exp(2 * of_real pi * \ * (1/n)) \ U" by blast+ then obtain y0 y1 where "y0 \ T" and y0: "(y0 - \) * g y0 = \ * exp(2 * of_real pi * \ * (0/n))" and "y1 \ T" and y1: "(y1 - \) * g y1 = \ * exp(2 * of_real pi * \ * (1/n))" by (auto simp: U_def) then have "y0 \ ball \ \" "y1 \ ball \ \" using Tsb by auto moreover have "y0 \ y1" using y0 y1 \\ > 0\ complex_root_unity_eq_1 [of n 1] \n > 0\ False by auto moreover have "T \ S" by (meson Tsb min.cobounded1 order_trans r subset_ball) ultimately have False using inj_onD [OF injf, of y0 y1] \y0 \ T\ \y1 \ T\ using fd [of y0] fd [of y1] complex_root_unity [of n 1] n_ne apply (simp add: y0 y1 power_mult_distrib) apply (force simp: algebra_simps) done then show ?thesis .. qed qed qed text\Hence a nice clean inverse function theorem\ lemma has_field_derivative_inverse_strong: fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" shows "\DERIV f x :> f'; f' \ 0; open S; x \ S; continuous_on S f; \z. z \ S \ g (f z) = z\ \ DERIV g (f x) :> inverse (f')" unfolding has_field_derivative_def by (rule has_derivative_inverse_strong [of S x f g]) auto lemma has_field_derivative_inverse_strong_x: fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" shows "\DERIV f (g y) :> f'; f' \ 0; open S; continuous_on S f; g y \ S; f(g y) = y; \z. z \ S \ g (f z) = z\ \ DERIV g y :> inverse (f')" unfolding has_field_derivative_def by (rule has_derivative_inverse_strong_x [of S g y f]) auto proposition holomorphic_has_inverse: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" obtains g where "g holomorphic_on (f ` S)" "\z. z \ S \ deriv f z * deriv g (f z) = 1" "\z. z \ S \ g(f z) = z" proof - have ofs: "open (f ` S)" by (rule open_mapping_thm3 [OF assms]) have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \ S" for z proof - have 1: "(f has_field_derivative deriv f z) (at z)" using DERIV_deriv_iff_field_differentiable \z \ S\ \open S\ holf holomorphic_on_imp_differentiable_at by blast have 2: "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast show ?thesis apply (rule has_field_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) apply (simp add: holf holomorphic_on_imp_continuous_on) by (simp add: injf the_inv_into_f_f) qed show ?thesis proof show "the_inv_into S f holomorphic_on f ` S" by (simp add: holomorphic_on_open ofs) (blast intro: *) next fix z assume "z \ S" have "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast then show "deriv f z * deriv (the_inv_into S f) (f z) = 1" using * [OF \z \ S\] by (simp add: DERIV_imp_deriv) next fix z assume "z \ S" show "the_inv_into S f (f z) = z" by (simp add: \z \ S\ injf the_inv_into_f_f) qed qed subsection\The Schwarz Lemma\ lemma Schwarz1: assumes holf: "f holomorphic_on S" and contf: "continuous_on (closure S) f" and S: "open S" "connected S" and boS: "bounded S" and "S \ {}" obtains w where "w \ frontier S" "\z. z \ closure S \ norm (f z) \ norm (f w)" proof - have connf: "continuous_on (closure S) (norm o f)" using contf continuous_on_compose continuous_on_norm_id by blast have coc: "compact (closure S)" by (simp add: \bounded S\ bounded_closure compact_eq_bounded_closed) then obtain x where x: "x \ closure S" and xmax: "\z. z \ closure S \ norm(f z) \ norm(f x)" apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]]) using \S \ {}\ apply auto done then show ?thesis proof (cases "x \ frontier S") case True then show ?thesis using that xmax by blast next case False then have "x \ S" using \open S\ frontier_def interior_eq x by auto then have "f constant_on S" apply (rule maximum_modulus_principle [OF holf S \open S\ order_refl]) using closure_subset apply (blast intro: xmax) done then have "f constant_on (closure S)" by (rule constant_on_closureI [OF _ contf]) then obtain c where c: "\x. x \ closure S \ f x = c" by (meson constant_on_def) obtain w where "w \ frontier S" by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV) then show ?thesis by (simp add: c frontier_def that) qed qed lemma Schwarz2: "\f holomorphic_on ball 0 r; 0 < s; ball w s \ ball 0 r; \z. norm (w-z) < s \ norm(f z) \ norm(f w)\ \ f constant_on ball 0 r" by (rule maximum_modulus_principle [where U = "ball w s" and \ = w]) (simp_all add: dist_norm) lemma Schwarz3: assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" obtains h where "h holomorphic_on (ball 0 r)" and "\z. norm z < r \ f z = z * (h z)" and "deriv f 0 = h 0" proof - define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z have d0: "deriv f 0 = h 0" by (simp add: h_def) moreover have "h holomorphic_on (ball 0 r)" by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) moreover have "norm z < r \ f z = z * h z" for z by (simp add: h_def) ultimately show ?thesis using that by blast qed proposition Schwarz_Lemma: assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" and no: "\z. norm z < 1 \ norm (f z) < 1" and \: "norm \ < 1" shows "norm (f \) \ norm \" and "norm(deriv f 0) \ 1" and "((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) \ \\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1" (is "?P \ ?Q") proof - obtain h where holh: "h holomorphic_on (ball 0 1)" and fz_eq: "\z. norm z < 1 \ f z = z * (h z)" and df0: "deriv f 0 = h 0" by (rule Schwarz3 [OF holf]) auto have noh_le: "norm (h z) \ 1" if z: "norm z < 1" for z proof - have "norm (h z) < a" if a: "1 < a" for a proof - have "max (inverse a) (norm z) < 1" using z a by (simp_all add: inverse_less_1_iff) then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" using Rats_dense_in_real by blast then have nzr: "norm z < r" and ira: "inverse r < a" using z a less_imp_inverse_less by force+ then have "0 < r" by (meson norm_not_less_zero not_le order.strict_trans2) have holh': "h holomorphic_on ball 0 r" by (meson holh \r < 1\ holomorphic_on_subset less_eq_real_def subset_ball) have conth': "continuous_on (cball 0 r) h" by (meson \r < 1\ dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) obtain w where w: "norm w = r" and lenw: "\z. norm z < r \ norm(h z) \ norm(h w)" apply (rule Schwarz1 [OF holh']) using conth' \0 < r\ by auto have "h w = f w / w" using fz_eq \r < 1\ nzr w by auto then have "cmod (h z) < inverse r" by (metis \0 < r\ \r < 1\ divide_strict_right_mono inverse_eq_divide le_less_trans lenw no norm_divide nzr w) then show ?thesis using ira by linarith qed then show "norm (h z) \ 1" using not_le by blast qed show "cmod (f \) \ cmod \" proof (cases "\ = 0") case True then show ?thesis by auto next case False then show ?thesis by (simp add: noh_le fz_eq \ mult_left_le norm_mult) qed show no_df0: "norm(deriv f 0) \ 1" by (simp add: \\z. cmod z < 1 \ cmod (h z) \ 1\ df0) show "?Q" if "?P" using that proof assume "\z. cmod z < 1 \ z \ 0 \ cmod (f z) = cmod z" then obtain \ where \: "cmod \ < 1" "\ \ 0" "cmod (f \) = cmod \" by blast then have [simp]: "norm (h \) = 1" by (simp add: fz_eq norm_mult) have "ball \ (1 - cmod \) \ ball 0 1" by (simp add: ball_subset_ball_iff) moreover have "\z. cmod (\ - z) < 1 - cmod \ \ cmod (h z) \ cmod (h \)" apply (simp add: algebra_simps) by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4) ultimately obtain c where c: "\z. norm z < 1 \ h z = c" using Schwarz2 [OF holh, of "1 - norm \" \, unfolded constant_on_def] \ by auto then have "norm c = 1" using \ by force with c show ?thesis using fz_eq by auto next assume [simp]: "cmod (deriv f 0) = 1" then obtain c where c: "\z. norm z < 1 \ h z = c" using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le by auto moreover have "norm c = 1" using df0 c by auto ultimately show ?thesis using fz_eq by auto qed qed corollary Schwarz_Lemma': assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" and no: "\z. norm z < 1 \ norm (f z) < 1" shows "((\\. norm \ < 1 \ norm (f \) \ norm \) \ norm(deriv f 0) \ 1) \ (((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) \ (\\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1))" using Schwarz_Lemma [OF assms] by (metis (no_types) norm_eq_zero zero_less_one) subsection\The Schwarz reflection principle\ lemma hol_pal_lem0: assumes "d \ a \ k" "k \ d \ b" obtains c where "c \ closed_segment a b" "d \ c = k" "\z. z \ closed_segment a c \ d \ z \ k" "\z. z \ closed_segment c b \ k \ d \ z" proof - obtain c where cin: "c \ closed_segment a b" and keq: "k = d \ c" using connected_ivt_hyperplane [of "closed_segment a b" a b d k] by (auto simp: assms) have "closed_segment a c \ {z. d \ z \ k}" "closed_segment c b \ {z. k \ d \ z}" unfolding segment_convex_hull using assms keq by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal) then show ?thesis using cin that by fastforce qed lemma hol_pal_lem1: assumes "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" "d \ 0" and lek: "d \ a \ k" "d \ b \ k" "d \ c \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof - have "interior (convex hull {a, b, c}) \ interior(S \ {x. d \ x \ k})" apply (rule interior_mono) apply (rule hull_minimal) apply (simp add: abc lek) apply (rule convex_Int [OF \convex S\ convex_halfspace_le]) done also have "... \ {z \ S. d \ z < k}" by (force simp: interior_open [OF \open S\] \d \ 0\) finally have *: "interior (convex hull {a, b, c}) \ {z \ S. d \ z < k}" . have "continuous_on (convex hull {a,b,c}) f" using \convex S\ contf abc continuous_on_subset subset_hull by fastforce moreover have "f holomorphic_on interior (convex hull {a,b,c})" by (rule holomorphic_on_subset [OF holf1 *]) ultimately show ?thesis using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3 by blast qed lemma hol_pal_lem2: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and lek: "d \ a \ k" "d \ b \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ c \ k") case True show ?thesis by (rule hol_pal_lem1 [OF S abc \d \ 0\ lek True holf1 contf]) next case False then have "d \ c > k" by force obtain a' where a': "a' \ closed_segment b c" and "d \ a' = k" and ba': "\z. z \ closed_segment b a' \ d \ z \ k" and a'c: "\z. z \ closed_segment a' c \ k \ d \ z" apply (rule hol_pal_lem0 [of d b k c, OF \d \ b \ k\]) using False by auto obtain b' where b': "b' \ closed_segment a c" and "d \ b' = k" and ab': "\z. z \ closed_segment a b' \ d \ z \ k" and b'c: "\z. z \ closed_segment b' c \ k \ d \ z" apply (rule hol_pal_lem0 [of d a k c, OF \d \ a \ k\]) using False by auto have a'b': "a' \ S \ b' \ S" using a' abc b' convex_contains_segment \convex S\ by auto have "continuous_on (closed_segment c a) f" by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) then have 1: "contour_integral (linepath c a) f = contour_integral (linepath c b') f + contour_integral (linepath b' a) f" apply (rule contour_integral_split_linepath) using b' by (simp add: closed_segment_commute) have "continuous_on (closed_segment b c) f" by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) then have 2: "contour_integral (linepath b c) f = contour_integral (linepath b a') f + contour_integral (linepath a' c) f" by (rule contour_integral_split_linepath [OF _ a']) have 3: "contour_integral (reversepath (linepath b' a')) f = - contour_integral (linepath b' a') f" by (rule contour_integral_reversepath [OF valid_path_linepath]) have fcd_le: "f field_differentiable at x" if "x \ interior S \ x \ interior {x. d \ x \ k}" for x proof - have "f holomorphic_on S \ {c. d \ c < k}" by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1) then have "\C D. x \ interior C \ interior D \ f holomorphic_on interior C \ interior D" using that by (metis Collect_mem_eq Int_Collect \d \ 0\ interior_halfspace_le interior_open \open S\) then show "f field_differentiable at x" by (metis at_within_interior holomorphic_on_def interior_Int interior_interior) qed have ab_le: "\x. x \ closed_segment a b \ d \ x \ k" proof - fix x :: complex assume "x \ closed_segment a b" then have "\C. x \ C \ b \ C \ a \ C \ \ convex C" by (meson contra_subsetD convex_contains_segment) then show "d \ x \ k" by (metis lek convex_halfspace_le mem_Collect_eq) qed have "continuous_on (S \ {x. d \ x \ k}) f" using contf by (simp add: continuous_on_subset) then have "(f has_contour_integral 0) (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)" apply (rule Cauchy_theorem_convex [where K = "{}"]) apply (simp_all add: path_image_join convex_Int convex_halfspace_le \convex S\ fcd_le ab_le closed_segment_subset abc a'b' ba') by (metis \d \ a' = k\ \d \ b' = k\ convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl) then have 4: "contour_integral (linepath a b) f + contour_integral (linepath b a') f + contour_integral (linepath a' b') f + contour_integral (linepath b' a) f = 0" by (rule has_chain_integral_chain_integral4) have fcd_ge: "f field_differentiable at x" if "x \ interior S \ x \ interior {x. k \ d \ x}" for x proof - have f2: "f holomorphic_on S \ {c. k < d \ c}" by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2) have f3: "interior S = S" by (simp add: interior_open \open S\) then have "x \ S \ interior {c. k \ d \ c}" using that by simp then show "f field_differentiable at x" using f3 f2 unfolding holomorphic_on_def by (metis (no_types) \d \ 0\ at_within_interior interior_Int interior_halfspace_ge interior_interior) qed have "continuous_on (S \ {x. k \ d \ x}) f" using contf by (simp add: continuous_on_subset) then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')" apply (rule Cauchy_theorem_convex [where K = "{}"]) apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \convex S\ fcd_ge closed_segment_subset abc a'b' a'c) by (metis \d \ a' = k\ b'c closed_segment_commute convex_contains_segment convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl) then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0" by (rule has_chain_integral_chain_integral3) show ?thesis using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath) qed lemma hol_pal_lem3: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and lek: "d \ a \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ b \ k") case True show ?thesis by (rule hol_pal_lem2 [OF S abc \d \ 0\ lek True holf1 holf2 contf]) next case False show ?thesis proof (cases "d \ c \ k") case True have "contour_integral (linepath c a) f + contour_integral (linepath a b) f + contour_integral (linepath b c) f = 0" by (rule hol_pal_lem2 [OF S \c \ S\ \a \ S\ \b \ S\ \d \ 0\ \d \ c \ k\ lek holf1 holf2 contf]) then show ?thesis by (simp add: algebra_simps) next case False have "contour_integral (linepath b c) f + contour_integral (linepath c a) f + contour_integral (linepath a b) f = 0" apply (rule hol_pal_lem2 [OF S \b \ S\ \c \ S\ \a \ S\, of "-d" "-k"]) using \d \ 0\ \\ d \ b \ k\ False by (simp_all add: holf1 holf2 contf) then show ?thesis by (simp add: algebra_simps) qed qed lemma hol_pal_lem4: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ a \ k") case True show ?thesis by (rule hol_pal_lem3 [OF S abc \d \ 0\ True holf1 holf2 contf]) next case False show ?thesis apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"]) using \d \ 0\ False by (simp_all add: holf1 holf2 contf) qed lemma holomorphic_on_paste_across_line: assumes S: "open S" and "d \ 0" and holf1: "f holomorphic_on (S \ {z. d \ z < k})" and holf2: "f holomorphic_on (S \ {z. k < d \ z})" and contf: "continuous_on S f" shows "f holomorphic_on S" proof - have *: "\t. open t \ p \ t \ continuous_on t f \ (\a b c. convex hull {a, b, c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" if "p \ S" for p proof - obtain e where "e>0" and e: "ball p e \ S" using \p \ S\ openE S by blast then have "continuous_on (ball p e) f" using contf continuous_on_subset by blast moreover have "f holomorphic_on {z. dist p z < e \ d \ z < k}" apply (rule holomorphic_on_subset [OF holf1]) using e by auto moreover have "f holomorphic_on {z. dist p z < e \ k < d \ z}" apply (rule holomorphic_on_subset [OF holf2]) using e by auto ultimately show ?thesis apply (rule_tac x="ball p e" in exI) using \e > 0\ e \d \ 0\ apply (simp add:, clarify) apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k]) apply (auto simp: subset_hull) done qed show ?thesis by (blast intro: * Morera_local_triangle analytic_imp_holomorphic) qed proposition Schwarz_reflection: assumes "open S" and cnjs: "cnj ` S \ S" and holf: "f holomorphic_on (S \ {z. 0 < Im z})" and contf: "continuous_on (S \ {z. 0 \ Im z}) f" and f: "\z. \z \ S; z \ \\ \ (f z) \ \" shows "(\z. if 0 \ Im z then f z else cnj(f(cnj z))) holomorphic_on S" proof - have 1: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. 0 < Im z})" by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf]) have cont_cfc: "continuous_on (S \ {z. Im z \ 0}) (cnj o f o cnj)" apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) using cnjs apply auto done have "cnj \ f \ cnj field_differentiable at x within S \ {z. Im z < 0}" if "x \ S" "Im x < 0" "f field_differentiable at (cnj x) within S \ {z. 0 < Im z}" for x using that apply (simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm, clarify) apply (rule_tac x="cnj f'" in exI) apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify) apply (drule_tac x="cnj xa" in bspec) using cnjs apply force apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj) done then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \ {z. Im z < 0})" using holf cnjs by (force simp: holomorphic_on_def) have 2: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. Im z < 0})" apply (rule iffD1 [OF holomorphic_cong [OF refl]]) using hol_cfc by auto have [simp]: "(S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0}) = S" by force have "continuous_on ((S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0})) (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" apply (rule continuous_on_cases_local) using cont_cfc contf apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge) using f Reals_cnj_iff complex_is_Real_iff apply auto done then have 3: "continuous_on S (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" by force show ?thesis apply (rule holomorphic_on_paste_across_line [OF \open S\, of "- \" _ 0]) using 1 2 3 apply auto done qed subsection\Bloch's theorem\ lemma Bloch_lemma_0: assumes holf: "f holomorphic_on cball 0 r" and "0 < r" and [simp]: "f 0 = 0" and le: "\z. norm z < r \ norm(deriv f z) \ 2 * norm(deriv f 0)" shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \ f ` ball 0 r" proof - have "sqrt 2 < 3/2" by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp show ?thesis proof (cases "deriv f 0 = 0") case True then show ?thesis by simp next case False define C where "C = 2 * norm(deriv f 0)" have "0 < C" using False by (simp add: C_def) have holf': "f holomorphic_on ball 0 r" using holf using ball_subset_cball holomorphic_on_subset by blast then have holdf': "deriv f holomorphic_on ball 0 r" by (rule holomorphic_deriv [OF _ open_ball]) have "Le1": "norm(deriv f z - deriv f 0) \ norm z / (r - norm z) * C" if "norm z < r" for z proof - have T1: "norm(deriv f z - deriv f 0) \ norm z / (R - norm z) * C" if R: "norm z < R" "R < r" for R proof - have "0 < R" using R by (metis less_trans norm_zero zero_less_norm_iff) have df_le: "\x. norm x < r \ norm (deriv f x) \ C" using le by (simp add: C_def) have hol_df: "deriv f holomorphic_on cball 0 R" apply (rule holomorphic_on_subset) using R holdf' by auto have *: "((\w. deriv f w / (w - z)) has_contour_integral 2 * pi * \ * deriv f z) (circlepath 0 R)" if "norm z < R" for z using \0 < R\ that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"] by (force simp: winding_number_circlepath) have **: "((\x. deriv f x / (x - z) - deriv f x / x) has_contour_integral of_real (2 * pi) * \ * (deriv f z - deriv f 0)) (circlepath 0 R)" using has_contour_integral_diff [OF * [of z] * [of 0]] \0 < R\ that by (simp add: algebra_simps) have [simp]: "\x. norm x = R \ x \ z" using that(1) by blast have "norm (deriv f x / (x - z) - deriv f x / x) \ C * norm z / (R * (R - norm z))" if "norm x = R" for x proof - have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) = norm (deriv f x) * norm z" by (simp add: norm_mult right_diff_distrib') show ?thesis using \0 < R\ \0 < C\ R that apply (simp add: norm_mult norm_divide divide_simps) using df_le norm_triangle_ineq2 \0 < C\ apply (auto intro!: mult_mono) done qed then show ?thesis using has_contour_integral_bound_circlepath [OF **, of "C * norm z/(R*(R - norm z))"] \0 < R\ \0 < C\ R apply (simp add: norm_mult norm_divide) apply (simp add: divide_simps mult.commute) done qed obtain r' where r': "norm z < r'" "r' < r" using Rats_dense_in_real [of "norm z" r] \norm z < r\ by blast then have [simp]: "closure {r'<.. norm(f z)" if r: "norm z < r" for z proof - have 1: "\x. x \ ball 0 r \ ((\z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0) (at x within ball 0 r)" by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+ have 2: "closed_segment 0 z \ ball 0 r" by (metis \0 < r\ convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that) have 3: "(\t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" apply (rule integrable_on_cmult_right [where 'b=real, simplified]) apply (rule integrable_on_cdivide [where 'b=real, simplified]) apply (rule integrable_on_cmult_left [where 'b=real, simplified]) apply (rule ident_integrable_on) done have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm z * norm z * x * C / (r - norm z)" if x: "0 \ x" "x \ 1" for x proof - have [simp]: "x * norm z < r" using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero) have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \ norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C" apply (rule Le1) using r x \0 < r\ by simp also have "... \ norm (x *\<^sub>R z) / (r - norm z) * C" using r x \0 < r\ apply (simp add: field_split_simps) by (simp add: \0 < C\ mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono) finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm (x *\<^sub>R z) / (r - norm z) * C * norm z" by (rule mult_right_mono) simp with x show ?thesis by (simp add: algebra_simps) qed have le_norm: "abc \ norm d - e \ norm(f - d) \ e \ abc \ norm f" for abc d e and f::complex by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans) have "norm (integral {0..1} (\x. (deriv f (x *\<^sub>R z) - deriv f 0) * z)) \ integral {0..1} (\t. (norm z)\<^sup>2 * t / (r - norm z) * C)" apply (rule integral_norm_bound_integral) using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 apply (simp add: has_contour_integral_linepath has_integral_integrable_integral) apply (rule 3) apply (simp add: norm_mult power2_eq_square 4) done then have int_le: "norm (f z - deriv f 0 * z) \ (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))" using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) done show ?thesis apply (rule le_norm [OF _ int_le]) using \norm z < r\ apply (simp add: power2_eq_square divide_simps C_def norm_mult) proof - have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \ norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" by (simp add: algebra_simps) then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \ norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" by (simp add: algebra_simps) qed qed have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2) < 1" by (auto simp: sqrt2_less_2) have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f" apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) apply (subst closure_ball) using \0 < r\ mult_pos_pos sq201 apply (auto simp: cball_subset_cball_iff) done have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))" apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force) using \0 < r\ mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff) using False \0 < r\ centre_in_ball holf' holomorphic_nonconstant by blast have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) = ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))" by simp also have "... \ f ` ball 0 ((1 - sqrt 2 / 2) * r)" proof - have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \ norm (f z)" if "norm z = (1 - sqrt 2 / 2) * r" for z apply (rule order_trans [OF _ *]) using \0 < r\ apply (simp_all add: field_simps power2_eq_square that) apply (simp add: mult.assoc [symmetric]) done show ?thesis apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball]) using \0 < r\ sq201 3 apply simp_all using C_def \0 < C\ sq3 apply force done qed also have "... \ f ` ball 0 r" apply (rule image_subsetI [OF imageI], simp) apply (erule less_le_trans) using \0 < r\ apply (auto simp: field_simps) done finally show ?thesis . qed qed lemma Bloch_lemma: assumes holf: "f holomorphic_on cball a r" and "0 < r" and le: "\z. z \ ball a r \ norm(deriv f z) \ 2 * norm(deriv f a)" shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \ f ` ball a r" proof - have fz: "(\z. f (a + z)) = f o (\z. (a + z))" by (simp add: o_def) have hol0: "(\z. f (a + z)) holomorphic_on cball 0 r" unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+ then have [simp]: "\x. norm x < r \ (\z. f (a + z)) field_differentiable at x" by (metis open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel) have [simp]: "\z. norm z < r \ f field_differentiable at (a + z)" by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute) then have [simp]: "f field_differentiable at a" by (metis add.comm_neutral \0 < r\ norm_eq_zero) have hol1: "(\z. f (a + z) - f a) holomorphic_on cball 0 r" by (intro holomorphic_intros hol0) then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\z. f (a + z) - f a) 0)) \ (\z. f (a + z) - f a) ` ball 0 r" apply (rule Bloch_lemma_0) apply (simp_all add: \0 < r\) apply (simp add: fz complex_derivative_chain) apply (simp add: dist_norm le) done then show ?thesis apply clarify apply (drule_tac c="x - f a" in subsetD) apply (force simp: fz \0 < r\ dist_norm complex_derivative_chain field_differentiable_compose)+ done qed proposition Bloch_unit: assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" obtains b r where "1/12 < r" and "ball b r \ f ` (ball a 1)" proof - define r :: real where "r = 249/256" have "0 < r" "r < 1" by (auto simp: r_def) define g where "g z = deriv f z * of_real(r - norm(z - a))" for z have "deriv f holomorphic_on ball a 1" by (rule holomorphic_deriv [OF holf open_ball]) then have "continuous_on (ball a 1) (deriv f)" using holomorphic_on_imp_continuous_on by blast then have "continuous_on (cball a r) (deriv f)" by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \r < 1\) then have "continuous_on (cball a r) g" by (simp add: g_def continuous_intros) then have 1: "compact (g ` cball a r)" by (rule compact_continuous_image [OF _ compact_cball]) have 2: "g ` cball a r \ {}" using \r > 0\ by auto obtain p where pr: "p \ cball a r" and pge: "\y. y \ cball a r \ norm (g y) \ norm (g p)" using distance_attains_sup [OF 1 2, of 0] by force define t where "t = (r - norm(p - a)) / 2" have "norm (p - a) \ r" using pge [of a] \r > 0\ by (auto simp: g_def norm_mult) then have "norm (p - a) < r" using pr by (simp add: norm_minus_commute dist_norm) then have "0 < t" by (simp add: t_def) have cpt: "cball p t \ ball a r" using \0 < t\ by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \ norm (deriv f p)" if "y \ cball a r" for y proof - have [simp]: "norm (y - a) \ r" using that by (simp add: dist_norm norm_minus_commute) have "norm (g y) \ norm (g p)" using pge [OF that] by simp then have "norm (deriv f y) * abs (r - norm (y - a)) \ norm (deriv f p) * abs (r - norm (p - a))" by (simp only: dist_norm g_def norm_mult norm_of_real) with that \norm (p - a) < r\ show ?thesis by (simp add: dist_norm field_split_simps) qed have le_norm_dfp: "r / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [of a] \r > 0\ by auto have 1: "f holomorphic_on cball p t" apply (rule holomorphic_on_subset [OF holf]) using cpt \r < 1\ order_subst1 subset_ball by auto have 2: "norm (deriv f z) \ 2 * norm (deriv f p)" if "z \ ball p t" for z proof - have z: "z \ cball a r" by (meson ball_subset_cball subsetD cpt that) then have "norm(z - a) < r" by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [OF z] by simp with \norm (z - a) < r\ \norm (p - a) < r\ have "norm (deriv f z) \ (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" by (simp add: field_simps) also have "... \ 2 * norm (deriv f p)" apply (rule mult_right_mono) using that \norm (p - a) < r\ \norm(z - a) < r\ apply (simp_all add: field_simps t_def dist_norm [symmetric]) using dist_triangle3 [of z a p] by linarith finally show ?thesis . qed have sqrt2: "sqrt 2 < 2113/1494" by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" using sq3 sqrt2 by (auto simp: field_simps r_def) also have "... \ cmod (deriv f p) * (r - cmod (p - a))" using \norm (p - a) < r\ le_norm_dfp by (simp add: pos_divide_le_eq) finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" using pos_divide_less_eq half_gt_zero_iff sq3 by blast then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" using sq3 by (simp add: mult.commute t_def) have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball p t" by (rule Bloch_lemma [OF 1 \0 < t\ 2]) also have "... \ f ` ball a 1" apply (rule image_mono) apply (rule order_trans [OF ball_subset_cball]) apply (rule order_trans [OF cpt]) using \0 < t\ \r < 1\ apply (simp add: ball_subset_ball_iff dist_norm) done finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball a 1" . with ** show ?thesis by (rule that) qed theorem Bloch: assumes holf: "f holomorphic_on ball a r" and "0 < r" and r': "r' \ r * norm (deriv f a) / 12" obtains b where "ball b r' \ f ` (ball a r)" proof (cases "deriv f a = 0") case True with r' show ?thesis using ball_eq_empty that by fastforce next case False define C where "C = deriv f a" have "0 < norm C" using False by (simp add: C_def) have dfa: "f field_differentiable at a" apply (rule holomorphic_on_imp_differentiable_at [OF holf]) using \0 < r\ by auto have fo: "(\z. f (a + of_real r * z)) = f o (\z. (a + of_real r * z))" by (simp add: o_def) have holf': "f holomorphic_on (\z. a + complex_of_real r * z) ` ball 0 1" apply (rule holomorphic_on_subset [OF holf]) using \0 < r\ apply (force simp: dist_norm norm_mult) done have 1: "(\z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ using \0 < r\ by (simp add: C_def False) have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative (deriv f (a + of_real r * z) / C)) (at z)" if "norm z < 1" for z proof - have *: "((\x. f (a + of_real r * x)) has_field_derivative (deriv f (a + of_real r * z) * of_real r)) (at z)" apply (simp add: fo) apply (rule DERIV_chain [OF field_differentiable_derivI]) apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) using \0 < r\ apply (simp add: dist_norm norm_mult that) apply (rule derivative_eq_intros | simp)+ done show ?thesis apply (rule derivative_eq_intros * | simp)+ using \0 < r\ by (auto simp: C_def False) qed have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" apply (subst deriv_cdivide_right) apply (simp add: field_differentiable_def fo) apply (rule exI) apply (rule DERIV_chain [OF field_differentiable_derivI]) apply (simp add: dfa) apply (rule derivative_eq_intros | simp add: C_def False fo)+ using \0 < r\ apply (simp add: C_def False fo) apply (simp add: derivative_intros dfa complex_derivative_chain) done have sb1: "(*) (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 \ f ` ball a r" using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) have sb2: "ball (C * r * b) r' \ (*) (C * r) ` ball b t" if "1 / 12 < t" for b t proof - have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right by auto show ?thesis apply clarify apply (rule_tac x="x / (C * r)" in image_eqI) using \0 < r\ apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) apply (erule less_le_trans) apply (rule order_trans [OF r' *]) done qed show ?thesis apply (rule Bloch_unit [OF 1 2]) apply (rename_tac t) apply (rule_tac b="(C * of_real r) * b" in that) apply (drule image_mono [where f = "\z. (C * of_real r) * z"]) using sb1 sb2 apply force done qed corollary Bloch_general: assumes holf: "f holomorphic_on s" and "a \ s" and tle: "\z. z \ frontier s \ t \ dist a z" and rle: "r \ t * norm(deriv f a) / 12" obtains b where "ball b r \ f ` s" proof - consider "r \ 0" | "0 < t * norm(deriv f a) / 12" using rle by force then show ?thesis proof cases case 1 then show ?thesis by (simp add: ball_empty that) next case 2 show ?thesis proof (cases "deriv f a = 0") case True then show ?thesis using rle by (simp add: ball_empty that) next case False then have "t > 0" using 2 by (force simp: zero_less_mult_iff) have "\ ball a t \ s \ ball a t \ frontier s \ {}" apply (rule connected_Int_frontier [of "ball a t" s], simp_all) using \0 < t\ \a \ s\ centre_in_ball apply blast done with tle have *: "ball a t \ s" by fastforce then have 1: "f holomorphic_on ball a t" using holf using holomorphic_on_subset by blast show ?thesis apply (rule Bloch [OF 1 \t > 0\ rle]) apply (rule_tac b=b in that) using * apply force done qed qed qed subsection \Cauchy's residue theorem\ text\Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. Interactive Theorem Proving\ definition\<^marker>\tag important\ residue :: "(complex \ complex) \ complex \ complex" where "residue f z = (SOME int. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" lemma Eps_cong: assumes "\x. P x = Q x" shows "Eps P = Eps Q" using ext[of P Q, OF assms] by simp lemma residue_cong: assumes eq: "eventually (\z. f z = g z) (at z)" and "z = z'" shows "residue f z = residue g z'" proof - from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) let ?P = "\f c e. (\\>0. \ < e \ (f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \))" have "residue f z = residue g z" unfolding residue_def proof (rule Eps_cong) fix c :: complex have "\e>0. ?P g c e" if "\e>0. ?P f c e" and "eventually (\z. f z = g z) (at z)" for f g proof - from that(1) obtain e where e: "e > 0" "?P f c e" by blast from that(2) obtain e' where e': "e' > 0" "\z'. z' \ z \ dist z' z < e' \ f z' = g z'" unfolding eventually_at by blast have "?P g c (min e e')" proof (intro allI exI impI, goal_cases) case (1 \) hence "(f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \)" using e(2) by auto thus ?case proof (rule has_contour_integral_eq) fix z' assume "z' \ path_image (circlepath z \)" hence "dist z' z < e'" and "z' \ z" using 1 by (auto simp: dist_commute) with e'(2)[of z'] show "f z' = g z'" by simp qed qed moreover from e and e' have "min e e' > 0" by auto ultimately show ?thesis by blast qed from this[OF _ eq] and this[OF _ eq'] show "(\e>0. ?P f c e) \ (\e>0. ?P g c e)" by blast qed with assms show ?thesis by simp qed lemma contour_integral_circlepath_eq: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" and e2_cball:"cball z e2 \ s" shows "f contour_integrable_on circlepath z e1" "f contour_integrable_on circlepath z e2" "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" proof - define l where "l \ linepath (z+e2) (z+e1)" have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto have "e2>0" using \e1>0\ \e1\e2\ by auto have zl_img:"z\path_image l" proof assume "z \ path_image l" then have "e2 \ cmod (e2 - e1)" using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \e1>0\ \e2>0\ unfolding l_def by (auto simp add:closed_segment_commute) thus False using \e2>0\ \e1>0\ \e1\e2\ apply (subst (asm) norm_of_real) by auto qed define g where "g \ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" proof - show "f contour_integrable_on circlepath z e2" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e2>0\ e2_cball by auto show "f contour_integrable_on (circlepath z e1)" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e1>0\ \e1\e2\ e2_cball by auto qed have [simp]:"f contour_integrable_on l" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) hence "closed_segment (z + e2) (z + e1) \ s - {z}" using zl_img e2_cball unfolding l_def by auto then show "f contour_integrable_on l" unfolding l_def apply (intro contour_integrable_continuous_linepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) by auto qed let ?ig="\g. contour_integral g f" have "(f has_contour_integral 0) g" proof (rule Cauchy_theorem_global[OF _ f_holo]) show "open (s - {z})" using \open s\ by auto show "valid_path g" unfolding g_def l_def by auto show "pathfinish g = pathstart g" unfolding g_def l_def by auto next have path_img:"path_image g \ cball z e2" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) moreover have "sphere z \e1\ \ cball z e2" using \e2>0\ \e1\e2\ \e1>0\ by auto ultimately show ?thesis unfolding g_def l_def using \e2>0\ by (simp add: path_image_join closed_segment_commute) qed show "path_image g \ s - {z}" proof - have "z\path_image g" using zl_img unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) moreover note \cball z e2 \ s\ and path_img ultimately show ?thesis by auto qed show "winding_number g w = 0" when"w \ s - {z}" for w proof - have "winding_number g w = 0" when "w\s" using that e2_cball apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) by (auto simp add:g_def l_def) moreover have "winding_number g z=0" proof - let ?Wz="\g. winding_number g z" have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + ?Wz (reversepath l)" using \e2>0\ \e1>0\ zl_img unfolding g_def l_def by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" using zl_img apply (subst (2) winding_number_reversepath) by (auto simp add:l_def closed_segment_commute) also have "... = 0" proof - have "?Wz (circlepath z e2) = 1" using \e2>0\ by (auto intro: winding_number_circlepath_centre) moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \e1>0\ apply (subst winding_number_reversepath) by (auto intro: winding_number_circlepath_centre) ultimately show ?thesis by auto qed finally show ?thesis . qed ultimately show ?thesis using that by auto qed qed then have "0 = ?ig g" using contour_integral_unique by simp also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + ?ig (reversepath l)" unfolding g_def by (auto simp add:contour_integrable_reversepath_eq) also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" by (auto simp add:contour_integral_reversepath) finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" by simp qed lemma base_residue: assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" and r_cball:"cball z r \ s" shows "(f has_contour_integral 2 * pi * \ * (residue f z)) (circlepath z r)" proof - obtain e where "e>0" and e_cball:"cball z e \ s" using open_contains_cball[of s] \open s\ \z\s\ by auto define c where "c \ 2 * pi * \" define i where "i \ contour_integral (circlepath z e) f / c" have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ proof - have "contour_integral (circlepath z e) f = contour_integral (circlepath z \) f" "f contour_integrable_on circlepath z \" "f contour_integrable_on circlepath z e" using \\ by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ then show ?thesis unfolding i_def c_def by (auto intro:has_contour_integral_integral) qed then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" unfolding residue_def c_def apply (rule_tac someI[of _ i],intro exI[where x=e]) by (auto simp add:\e>0\ c_def) then obtain e' where "e'>0" and e'_def:"\\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" by auto let ?int="\e. contour_integral (circlepath z e) f" define \ where "\ \ Min {r,e'} / 2" have "\>0" "\\r" "\r>0\ \e'>0\ unfolding \_def by auto have "(f has_contour_integral c * (residue f z)) (circlepath z \)" using e'_def[rule_format,OF \\>0\ \\] . then show ?thesis unfolding c_def using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \" "circlepath z r"]) qed lemma residue_holo: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s" shows "residue f z = 0" proof - define c where "c \ 2 * pi * \" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f has_contour_integral c*residue f z) (circlepath z e)" using f_holo by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) moreover have "(f has_contour_integral 0) (circlepath z e)" using f_holo e_cball \e>0\ by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) ultimately have "c*residue f z =0" using has_contour_integral_unique by blast thus ?thesis unfolding c_def by auto qed lemma residue_const:"residue (\_. c) z = 0" by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) lemma residue_add: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z + g z) z= residue f z + residue g z" proof - define c where "c \ 2 * pi * \" define fg where "fg \ (\z. f z+g z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(fg has_contour_integral c * residue fg z) (circlepath z e)" unfolding fg_def using f_holo g_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro:holomorphic_intros) moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" unfolding fg_def using f_holo g_holo by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) ultimately have "c*(residue f z + residue g z) = c * residue fg z" using has_contour_integral_unique by (auto simp add:distrib_left) thus ?thesis unfolding fg_def by (auto simp add:c_def) qed lemma residue_lmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. c * (f z)) z= c * residue f z" proof (cases "c=0") case True thus ?thesis using residue_const by auto next case False define c' where "c' \ 2 * pi * \" define f' where "f' \ (\z. c * (f z))" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) by (auto intro:holomorphic_intros) moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" unfolding f'_def using f_holo by (auto intro: has_contour_integral_lmul base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) ultimately have "c' * residue f' z = c * (c' * residue f z)" using has_contour_integral_unique by auto thus ?thesis unfolding f'_def c'_def using False by (auto simp add:field_simps) qed lemma residue_rmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) * c) z= residue f z * c" using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) lemma residue_div: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) / c) z= residue f z / c " using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) lemma residue_neg: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. - (f z)) z= - residue f z" using residue_lmul[OF assms,of "-1"] by auto lemma residue_diff: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z - g z) z= residue f z - residue g z" using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] by (auto intro:holomorphic_intros g_holo) lemma residue_simple: assumes "open s" "z\s" and f_holo:"f holomorphic_on s" shows "residue (\w. f w / (w - z)) z = f z" proof - define c where "c \ 2 * pi * \" define f' where "f' \ \w. f w / (w - z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c * f z) (circlepath z e)" unfolding f'_def c_def using \e>0\ f_holo e_cball by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro!:holomorphic_intros) ultimately have "c * f z = c * residue f' z" using has_contour_integral_unique by blast thus ?thesis unfolding c_def f'_def by auto qed lemma residue_simple': assumes s: "open s" "z \ s" and holo: "f holomorphic_on (s - {z})" and lim: "((\w. f w * (w - z)) \ c) (at z)" shows "residue f z = c" proof - define g where "g = (\w. if w = z then c else f w * (w - z))" from holo have "(\w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P") by (force intro: holomorphic_intros) also have "?P \ g holomorphic_on (s - {z})" by (intro holomorphic_cong refl) (simp_all add: g_def) finally have *: "g holomorphic_on (s - {z})" . note lim also have "(\w. f w * (w - z)) \z\ c \ g \z\ g z" by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter) finally have **: "g \z\ g z" . have g_holo: "g holomorphic_on s" by (rule no_isolated_singularity'[where K = "{z}"]) (insert assms * **, simp_all add: at_within_open_NO_MATCH) from s and this have "residue (\w. g w / (w - z)) z = g z" by (rule residue_simple) also have "\\<^sub>F za in at z. g za / (za - z) = f za" unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def) hence "residue (\w. g w / (w - z)) z = residue f z" by (intro residue_cong refl) finally show ?thesis by (simp add: g_def) qed lemma residue_holomorphic_over_power: assumes "open A" "z0 \ A" "f holomorphic_on A" shows "residue (\z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n" proof - let ?f = "\z. f z / (z - z0) ^ Suc n" from assms(1,2) obtain r where r: "r > 0" "cball z0 r \ A" by (auto simp: open_contains_cball) have "(?f has_contour_integral 2 * pi * \ * residue ?f z0) (circlepath z0 r)" using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros) moreover have "(?f has_contour_integral 2 * pi * \ / fact n * (deriv ^^ n) f z0) (circlepath z0 r)" using assms r by (intro Cauchy_has_contour_integral_higher_derivative_circlepath) (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on) ultimately have "2 * pi * \ * residue ?f z0 = 2 * pi * \ / fact n * (deriv ^^ n) f z0" by (rule has_contour_integral_unique) thus ?thesis by (simp add: field_simps) qed lemma residue_holomorphic_over_power': assumes "open A" "0 \ A" "f holomorphic_on A" shows "residue (\z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" using residue_holomorphic_over_power[OF assms] by simp lemma get_integrable_path: assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\s-pts" "b\s-pts" obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" "path_image g \ s-pts" "f contour_integrable_on g" using assms proof (induct arbitrary:s thesis a rule:finite_induct[OF \finite pts\]) case 1 obtain g where "valid_path g" "path_image g \ s" "pathstart g = a" "pathfinish g = b" using connected_open_polynomial_connected[OF \open s\,of a b ] \connected (s - {})\ valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto moreover have "f contour_integrable_on g" using contour_integrable_holomorphic_simple[OF _ \open s\ \valid_path g\ \path_image g \ s\,of f] \f holomorphic_on s - {}\ by auto ultimately show ?case using "1"(1)[of g] by auto next case idt:(2 p pts) obtain e where "e>0" and e:"\w\ball a e. w \ s \ (w \ a \ w \ insert p pts)" using finite_ball_avoid[OF \open s\ \finite (insert p pts)\, of a] \a \ s - insert p pts\ by auto define a' where "a' \ a+e/2" have "a'\s-{p} -pts" using e[rule_format,of "a+e/2"] \e>0\ by (auto simp add:dist_complex_def a'_def) then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" "path_image g' \ s - {p} - pts" "f contour_integrable_on g'" using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) by (metis Diff_insert2 open_delete) define g where "g \ linepath a a' +++ g'" have "valid_path g" unfolding g_def by (auto intro: valid_path_join) moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto moreover have "path_image g \ s - insert p pts" unfolding g_def proof (rule subset_path_image_join) have "closed_segment a a' \ ball a e" using \e>0\ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) then show "path_image (linepath a a') \ s - insert p pts" using e idt(9) by auto next show "path_image g' \ s - insert p pts" using g'(4) by blast qed moreover have "f contour_integrable_on g" proof - have "closed_segment a a' \ ball a e" using \e>0\ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) then have "continuous_on (closed_segment a a') f" using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] apply (elim continuous_on_subset) by auto then have "f contour_integrable_on linepath a a'" using contour_integrable_continuous_linepath by auto then show ?thesis unfolding g_def apply (rule contour_integrable_joinI) by (auto simp add: \e>0\) qed ultimately show ?case using idt.prems(1)[of g] by auto qed lemma Cauchy_theorem_aux: assumes "open s" "connected (s-pts)" "finite pts" "pts \ s" "f holomorphic_on s-pts" "valid_path g" "pathfinish g = pathstart g" "path_image g \ s-pts" "\z. (z \ s) \ winding_number g z = 0" "\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" using assms proof (induct arbitrary:s g rule:finite_induct[OF \finite pts\]) case 1 then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique) next case (2 p pts) note fin[simp] = \finite (insert p pts)\ and connected = \connected (s - insert p pts)\ and valid[simp] = \valid_path g\ and g_loop[simp] = \pathfinish g = pathstart g\ and holo[simp]= \f holomorphic_on s - insert p pts\ and path_img = \path_image g \ s - insert p pts\ and winding = \\z. z \ s \ winding_number g z = 0\ and h = \\pa\s. 0 < h pa \ (\w\cball pa (h pa). w \ s \ (w \ pa \ w \ insert p pts))\ have "h p>0" and "p\s" and h_p: "\w\cball p (h p). w \ s \ (w \ p \ w \ insert p pts)" using h \insert p pts \ s\ by auto obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" "path_image pg \ s-insert p pts" "f contour_integrable_on pg" proof - have "p + h p\cball p (h p)" using h[rule_format,of p] by (simp add: \p \ s\ dist_norm) then have "p + h p \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ by fastforce moreover have "pathstart g \ s - insert p pts " using path_img by auto ultimately show ?thesis using get_integrable_path[OF \open s\ connected fin holo,of "pathstart g" "p+h p"] that by blast qed obtain n::int where "n=winding_number g p" using integer_winding_number[OF _ g_loop,of p] valid path_img by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) define p_circ where "p_circ \ circlepath p (h p)" define p_circ_pt where "p_circ_pt \ linepath (p+h p) (p+h p)" define n_circ where "n_circ \ \n. ((+++) p_circ ^^ n) p_circ_pt" define cp where "cp \ if n\0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" have n_circ:"valid_path (n_circ k)" "winding_number (n_circ k) p = k" "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" "path_image (n_circ k) = (if k=0 then {p + h p} else sphere p (h p))" "p \ path_image (n_circ k)" "\p'. p'\s - pts \ winding_number (n_circ k) p'=0 \ p'\path_image (n_circ k)" "f contour_integrable_on (n_circ k)" "contour_integral (n_circ k) f = k * contour_integral p_circ f" for k proof (induct k) case 0 show "valid_path (n_circ 0)" and "path_image (n_circ 0) = (if 0=0 then {p + h p} else sphere p (h p))" and "winding_number (n_circ 0) p = of_nat 0" and "pathstart (n_circ 0) = p + h p" and "pathfinish (n_circ 0) = p + h p" and "p \ path_image (n_circ 0)" unfolding n_circ_def p_circ_pt_def using \h p > 0\ by (auto simp add: dist_norm) show "winding_number (n_circ 0) p'=0 \ p'\path_image (n_circ 0)" when "p'\s- pts" for p' unfolding n_circ_def p_circ_pt_def apply (auto intro!:winding_number_trivial) by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+ show "f contour_integrable_on (n_circ 0)" unfolding n_circ_def p_circ_pt_def by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing) show "contour_integral (n_circ 0) f = of_nat 0 * contour_integral p_circ f" unfolding n_circ_def p_circ_pt_def by auto next case (Suc k) have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto have pcirc:"p \ path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" using Suc(3) unfolding p_circ_def using \h p > 0\ by (auto simp add: p_circ_def) have pcirc_image:"path_image p_circ \ s - insert p pts" proof - have "path_image p_circ \ cball p (h p)" using \0 < h p\ p_circ_def by auto then show ?thesis using h_p pcirc(1) by auto qed have pcirc_integrable:"f contour_integrable_on p_circ" by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on holomorphic_on_subset[OF holo]) show "valid_path (n_circ (Suc k))" using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto show "path_image (n_circ (Suc k)) = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" proof - have "path_image p_circ = sphere p (h p)" unfolding p_circ_def using \0 < h p\ by auto then show ?thesis unfolding n_Suc using Suc.hyps(5) \h p>0\ by (auto simp add: path_image_join[OF pcirc(3)] dist_norm) qed then show "p \ path_image (n_circ (Suc k))" using \h p>0\ by auto show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" proof - have "winding_number p_circ p = 1" by (simp add: \h p > 0\ p_circ_def winding_number_circlepath_centre) moreover have "p \ path_image (n_circ k)" using Suc(5) \h p>0\ by auto then have "winding_number (p_circ +++ n_circ k) p = winding_number p_circ p + winding_number (n_circ k) p" using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc apply (intro winding_number_join) by auto ultimately show ?thesis using Suc(2) unfolding n_circ_def by auto qed show "pathstart (n_circ (Suc k)) = p + h p" by (simp add: n_circ_def p_circ_def) show "pathfinish (n_circ (Suc k)) = p + h p" using Suc(4) unfolding n_circ_def by auto show "winding_number (n_circ (Suc k)) p'=0 \ p'\path_image (n_circ (Suc k))" when "p'\s-pts" for p' proof - have " p' \ path_image p_circ" using \p \ s\ h p_circ_def that using pcirc_image by blast moreover have "p' \ path_image (n_circ k)" using Suc.hyps(7) that by blast moreover have "winding_number p_circ p' = 0" proof - have "path_image p_circ \ cball p (h p)" using h unfolding p_circ_def using \p \ s\ by fastforce moreover have "p'\cball p (h p)" using \p \ s\ h that "2.hyps"(2) by fastforce ultimately show ?thesis unfolding p_circ_def apply (intro winding_number_zero_outside) by auto qed ultimately show ?thesis unfolding n_Suc apply (subst winding_number_join) by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that]) qed show "f contour_integrable_on (n_circ (Suc k))" unfolding n_Suc by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]) show "contour_integral (n_circ (Suc k)) f = (Suc k) * contour_integral p_circ f" unfolding n_Suc by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] Suc(9) algebra_simps) qed have cp[simp]:"pathstart cp = p + h p" "pathfinish cp = p + h p" "valid_path cp" "path_image cp \ s - insert p pts" "winding_number cp p = - n" "\p'. p'\s - pts \ winding_number cp p'=0 \ p' \ path_image cp" "f contour_integrable_on cp" "contour_integral cp f = - n * contour_integral p_circ f" proof - show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp" using n_circ unfolding cp_def by auto next have "sphere p (h p) \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ by force moreover have "p + complex_of_real (h p) \ s - insert p pts" using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE) ultimately show "path_image cp \ s - insert p pts" unfolding cp_def using n_circ(5) by auto next show "winding_number cp p = - n" unfolding cp_def using winding_number_reversepath n_circ \h p>0\ by (auto simp: valid_path_imp_path) next show "winding_number cp p'=0 \ p' \ path_image cp" when "p'\s - pts" for p' unfolding cp_def apply (auto) apply (subst winding_number_reversepath) by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1)) next show "f contour_integrable_on cp" unfolding cp_def using contour_integrable_reversepath_eq n_circ(1,8) by auto next show "contour_integral cp f = - n * contour_integral p_circ f" unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9) by auto qed define g' where "g' \ g +++ pg +++ cp +++ (reversepath pg)" have "contour_integral g' f = (\p\pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \finite pts\ ]) show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) show "open (s - {p})" using \open s\ by auto show " pts \ s - {p}" using \insert p pts \ s\ \ p \ pts\ by blast show "f holomorphic_on s - {p} - pts" using holo \p \ pts\ by (metis Diff_insert2) show "valid_path g'" unfolding g'_def cp_def using n_circ valid pg g_loop by (auto intro!:valid_path_join ) show "pathfinish g' = pathstart g'" unfolding g'_def cp_def using pg(2) by simp show "path_image g' \ s - {p} - pts" proof - define s' where "s' \ s - {p} - pts" have s':"s' = s-insert p pts " unfolding s'_def by auto then show ?thesis using path_img pg(4) cp(4) unfolding g'_def apply (fold s'_def s') apply (intro subset_path_image_join) by auto qed note path_join_imp[simp] show "\z. z \ s - {p} \ winding_number g' z = 0" proof clarify fix z assume z:"z\s - {p}" have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z + winding_number (pg +++ cp +++ (reversepath pg)) z" proof (rule winding_number_join) show "path g" using \valid_path g\ by (simp add: valid_path_imp_path) show "z \ path_image g" using z path_img by auto show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp by (simp add: valid_path_imp_path) next have "path_image (pg +++ cp +++ reversepath pg) \ s - insert p pts" using pg(4) cp(4) by (auto simp:subset_path_image_join) then show "z \ path_image (pg +++ cp +++ reversepath pg)" using z by auto next show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto qed also have "... = winding_number g z + (winding_number pg z + winding_number (cp +++ (reversepath pg)) z)" proof (subst add_left_cancel,rule winding_number_join) show "path pg" and "path (cp +++ reversepath pg)" and "pathfinish pg = pathstart (cp +++ reversepath pg)" by (auto simp add: valid_path_imp_path) show "z \ path_image pg" using pg(4) z by blast show "z \ path_image (cp +++ reversepath pg)" using z by (metis Diff_iff \z \ path_image pg\ contra_subsetD cp(4) insertI1 not_in_path_image_join path_image_reversepath singletonD) qed also have "... = winding_number g z + (winding_number pg z + (winding_number cp z + winding_number (reversepath pg) z))" apply (auto intro!:winding_number_join simp: valid_path_imp_path) apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z) by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) also have "... = winding_number g z + winding_number cp z" apply (subst winding_number_reversepath) apply (auto simp: valid_path_imp_path) by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z) finally have "winding_number g' z = winding_number g z + winding_number cp z" unfolding g'_def . moreover have "winding_number g z + winding_number cp z = 0" using winding z \n=winding_number g p\ by auto ultimately show "winding_number g' z = 0" unfolding g'_def by auto qed show "\pa\s - {p}. 0 < h pa \ (\w\cball pa (h pa). w \ s - {p} \ (w \ pa \ w \ pts))" using h by fastforce qed moreover have "contour_integral g' f = contour_integral g f - winding_number g p * contour_integral p_circ f" proof - have "contour_integral g' f = contour_integral g f + contour_integral (pg +++ cp +++ reversepath pg) f" unfolding g'_def apply (subst contour_integral_join) by (auto simp add:open_Diff[OF \open s\,OF finite_imp_closed[OF fin]] intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral pg f + contour_integral (cp +++ reversepath pg) f" apply (subst contour_integral_join) by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral pg f + contour_integral cp f + contour_integral (reversepath pg) f" apply (subst contour_integral_join) by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral cp f" using contour_integral_reversepath by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f" using \n=winding_number g p\ by auto finally show ?thesis . qed moreover have "winding_number g' p' = winding_number g p'" when "p'\pts" for p' proof - have [simp]: "p' \ path_image g" "p' \ path_image pg" "p'\path_image cp" using "2.prems"(8) that apply blast apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that) by (meson DiffD2 cp(4) rev_subsetD subset_insertI that) have "winding_number g' p' = winding_number g p' + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def apply (subst winding_number_join) apply (simp_all add: valid_path_imp_path) apply (intro not_in_path_image_join) by auto also have "... = winding_number g p' + winding_number pg p' + winding_number (cp +++ reversepath pg) p'" apply (subst winding_number_join) apply (simp_all add: valid_path_imp_path) apply (intro not_in_path_image_join) by auto also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p' + winding_number (reversepath pg) p'" apply (subst winding_number_join) by (simp_all add: valid_path_imp_path) also have "... = winding_number g p' + winding_number cp p'" apply (subst winding_number_reversepath) by (simp_all add: valid_path_imp_path) also have "... = winding_number g p'" using that by auto finally show ?thesis . qed ultimately show ?case unfolding p_circ_def apply (subst (asm) sum.cong[OF refl, of pts _ "\p. winding_number g p * contour_integral (circlepath p (h p)) f"]) by (auto simp add:sum.insert[OF \finite pts\ \p\pts\] algebra_simps) qed lemma Cauchy_theorem_singularities: assumes "open s" "connected s" "finite pts" and holo:"f holomorphic_on s-pts" and "valid_path g" and loop:"pathfinish g = pathstart g" and "path_image g \ s-pts" and homo:"\z. (z \ s) \ winding_number g z = 0" and avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" (is "?L=?R") proof - define circ where "circ \ \p. winding_number g p * contour_integral (circlepath p (h p)) f" define pts1 where "pts1 \ pts \ s" define pts2 where "pts2 \ pts - pts1" have "pts=pts1 \ pts2" "pts1 \ pts2 = {}" "pts2 \ s={}" "pts1\s" unfolding pts1_def pts2_def by auto have "contour_integral g f = (\p\pts1. circ p)" unfolding circ_def proof (rule Cauchy_theorem_aux[OF \open s\ _ _ \pts1\s\ _ \valid_path g\ loop _ homo]) have "finite pts1" unfolding pts1_def using \finite pts\ by auto then show "connected (s - pts1)" using \open s\ \connected s\ connected_open_delete_finite[of s] by auto next show "finite pts1" using \pts = pts1 \ pts2\ assms(3) by auto show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def) show "path_image g \ s - pts1" using assms(7) pts1_def by auto show "\p\s. 0 < h p \ (\w\cball p (h p). w \ s \ (w \ p \ w \ pts1))" by (simp add: avoid pts1_def) qed moreover have "sum circ pts2=0" proof - have "winding_number g p=0" when "p\pts2" for p using \pts2 \ s={}\ that homo[rule_format,of p] by auto thus ?thesis unfolding circ_def apply (intro sum.neutral) by auto qed moreover have "?R=sum circ pts1 + sum circ pts2" unfolding circ_def using sum.union_disjoint[OF _ _ \pts1 \ pts2 = {}\] \finite pts\ \pts=pts1 \ pts2\ by blast ultimately show ?thesis apply (fold circ_def) by auto qed theorem Residue_theorem: fixes s pts::"complex set" and f::"complex \ complex" and g::"real \ complex" assumes "open s" "connected s" "finite pts" and holo:"f holomorphic_on s-pts" and "valid_path g" and loop:"pathfinish g = pathstart g" and "path_image g \ s-pts" and homo:"\z. (z \ s) \ winding_number g z = 0" shows "contour_integral g f = 2 * pi * \ *(\p\pts. winding_number g p * residue f p)" proof - define c where "c \ 2 * pi * \" obtain h where avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" using finite_cball_avoid[OF \open s\ \finite pts\] by metis have "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" using Cauchy_theorem_singularities[OF assms avoid] . also have "... = (\p\pts. c * winding_number g p * residue f p)" proof (intro sum.cong) show "pts = pts" by simp next fix x assume "x \ pts" show "winding_number g x * contour_integral (circlepath x (h x)) f = c * winding_number g x * residue f x" proof (cases "x\s") case False then have "winding_number g x=0" using homo by auto thus ?thesis by auto next case True have "contour_integral (circlepath x (h x)) f = c* residue f x" using \x\pts\ \finite pts\ avoid[rule_format,OF True] apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def]) by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \open s\ finite_imp_closed]) then show ?thesis by auto qed qed also have "... = c * (\p\pts. winding_number g p * residue f p)" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis unfolding c_def . qed subsection \Non-essential singular points\ definition\<^marker>\tag important\ is_pole :: "('a::topological_space \ 'b::real_normed_vector) \ 'a \ bool" where "is_pole f a = (LIM x (at a). f x :> at_infinity)" lemma is_pole_cong: assumes "eventually (\x. f x = g x) (at a)" "a=b" shows "is_pole f a \ is_pole g b" unfolding is_pole_def using assms by (intro filterlim_cong,auto) lemma is_pole_transform: assumes "is_pole f a" "eventually (\x. f x = g x) (at a)" "a=b" shows "is_pole g b" using is_pole_cong assms by auto lemma is_pole_tendsto: fixes f::"('a::topological_space \ 'b::real_normed_div_algebra)" shows "is_pole f x \ ((inverse o f) \ 0) (at x)" unfolding is_pole_def by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at) lemma is_pole_inverse_holomorphic: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and pole:"is_pole f z" and non_z:"\x\s-{z}. f x\0" shows "(\x. if x=z then 0 else inverse (f x)) holomorphic_on s" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" have "isCont g z" unfolding isCont_def using is_pole_tendsto[OF pole] apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \ f"]) by (simp_all add:g_def) moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def by (auto elim!:continuous_on_inverse simp add:non_z) hence "continuous_on (s-{z}) g" unfolding g_def apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"]) by auto ultimately have "continuous_on s g" using open_delete[OF \open s\] \open s\ by (auto simp add:continuous_on_eq_continuous_at) moreover have "(inverse o f) holomorphic_on (s-{z})" unfolding comp_def using f_holo by (auto elim!:holomorphic_on_inverse simp add:non_z) hence "g holomorphic_on (s-{z})" apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"]) by (auto simp add:g_def) ultimately show ?thesis unfolding g_def using \open s\ by (auto elim!: no_isolated_singularity) qed lemma not_is_pole_holomorphic: assumes "open A" "x \ A" "f holomorphic_on A" shows "\is_pole f x" proof - have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at) hence "f \x\ f x" by (simp add: isCont_def) thus "\is_pole f x" unfolding is_pole_def using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto qed lemma is_pole_inverse_power: "n > 0 \ is_pole (\z::complex. 1 / (z - a) ^ n) a" unfolding is_pole_def inverse_eq_divide [symmetric] by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros) (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros) lemma is_pole_inverse: "is_pole (\z::complex. 1 / (z - a)) a" using is_pole_inverse_power[of 1 a] by simp lemma is_pole_divide: fixes f :: "'a :: t2_space \ 'b :: real_normed_field" assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \ 0" shows "is_pole (\z. f z / g z) z" proof - have "filterlim (\z. f z * inverse (g z)) at_infinity (at z)" by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"] filterlim_compose[OF filterlim_inverse_at_infinity])+ (insert assms, auto simp: isCont_def) thus ?thesis by (simp add: field_split_simps is_pole_def) qed lemma is_pole_basic: assumes "f holomorphic_on A" "open A" "z \ A" "f z \ 0" "n > 0" shows "is_pole (\w. f w / (w - z) ^ n) z" proof (rule is_pole_divide) have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at) have "filterlim (\w. (w - z) ^ n) (nhds 0) (at z)" using assms by (auto intro!: tendsto_eq_intros) thus "filterlim (\w. (w - z) ^ n) (at 0) (at z)" by (intro filterlim_atI tendsto_eq_intros) (insert assms, auto simp: eventually_at_filter) qed fact+ lemma is_pole_basic': assumes "f holomorphic_on A" "open A" "0 \ A" "f 0 \ 0" "n > 0" shows "is_pole (\w. f w / w ^ n) 0" using is_pole_basic[of f A 0] assms by simp text \The proposition \<^term>\\x. ((f::complex\complex) \ x) (at z) \ is_pole f z\ can be interpreted as the complex function \<^term>\f\ has a non-essential singularity at \<^term>\z\ (i.e. the singularity is either removable or a pole).\ definition not_essential::"[complex \ complex, complex] \ bool" where "not_essential f z = (\x. f\z\x \ is_pole f z)" definition isolated_singularity_at::"[complex \ complex, complex] \ bool" where "isolated_singularity_at f z = (\r>0. f analytic_on ball z r-{z})" named_theorems singularity_intros "introduction rules for singularities" lemma holomorphic_factor_unique: fixes f::"complex \ complex" and z::complex and r::real and m n::int assumes "r>0" "g z\0" "h z\0" and asm:"\w\ball z r-{z}. f w = g w * (w-z) powr n \ g w\0 \ f w = h w * (w - z) powr m \ h w\0" and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" shows "n=m" proof - have [simp]:"at z within ball z r \ bot" using \r>0\ by (auto simp add:at_within_ball_bot_iff) have False when "n>m" proof - have "(h \ 0) (at z within ball z r)" proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (n - m) * g w"]) have "\w\ball z r-{z}. h w = (w-z)powr(n-m) * g w" using \n>m\ asm \r>0\ apply (auto simp add:field_simps powr_diff) by force then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) powr (n - m) * g x' = h x'" for x' by auto next define F where "F \ at z within ball z r" define f' where "f' \ \x. (x - z) powr (n-m)" have "f' z=0" using \n>m\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def continuous_def apply (subst Lim_ident_at) using \n>m\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) ultimately have "(f' \ 0) F" unfolding F_def by (simp add: continuous_within) moreover have "(g \ g z) F" using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \r>0\ unfolding F_def by auto ultimately show " ((\w. f' w * g w) \ 0) F" using tendsto_mult by fastforce qed moreover have "(h \ h z) (at z within ball z r)" using holomorphic_on_imp_continuous_on[OF h_holo] by (auto simp add:continuous_on_def \r>0\) ultimately have "h z=0" by (auto intro!: tendsto_unique) thus False using \h z\0\ by auto qed moreover have False when "m>n" proof - have "(g \ 0) (at z within ball z r)" proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (m - n) * h w"]) have "\w\ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \m>n\ asm apply (auto simp add:field_simps powr_diff) by force then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) powr (m - n) * h x' = g x'" for x' by auto next define F where "F \ at z within ball z r" define f' where "f' \\x. (x - z) powr (m-n)" have "f' z=0" using \m>n\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def continuous_def apply (subst Lim_ident_at) using \m>n\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) ultimately have "(f' \ 0) F" unfolding F_def by (simp add: continuous_within) moreover have "(h \ h z) F" using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \r>0\ unfolding F_def by auto ultimately show " ((\w. f' w * h w) \ 0) F" using tendsto_mult by fastforce qed moreover have "(g \ g z) (at z within ball z r)" using holomorphic_on_imp_continuous_on[OF g_holo] by (auto simp add:continuous_on_def \r>0\) ultimately have "g z=0" by (auto intro!: tendsto_unique) thus False using \g z\0\ by auto qed ultimately show "n=m" by fastforce qed lemma holomorphic_factor_puncture: assumes f_iso:"isolated_singularity_at f z" and "not_essential f z" \ \\<^term>\f\ has either a removable singularity or a pole at \<^term>\z\\ and non_zero:"\\<^sub>Fw in (at z). f w\0" \ \\<^term>\f\ will not be constantly zero in a neighbour of \<^term>\z\\ shows "\!n::int. \g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r-{z}. f w = g w * (w-z) powr n \ g w\0)" proof - define P where "P = (\f n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have imp_unique:"\!n::int. \g r. P f n g r" when "\n g r. P f n g r" proof (rule ex_ex1I[OF that]) fix n1 n2 :: int assume g1_asm:"\g1 r1. P f n1 g1 r1" and g2_asm:"\g2 r2. P f n2 g2 r2" define fac where "fac \ \n g r. \w\cball z r-{z}. f w = g w * (w - z) powr (of_int n) \ g w \ 0" obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\0" and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\0" and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def by auto define r where "r \ min r1 r2" have "r>0" using \r1>0\ \r2>0\ unfolding r_def by auto moreover have "\w\ball z r-{z}. f w = g1 w * (w-z) powr n1 \ g1 w\0 \ f w = g2 w * (w - z) powr n2 \ g2 w\0" using \fac n1 g1 r1\ \fac n2 g2 r2\ unfolding fac_def r_def by fastforce ultimately show "n1=n2" using g1_holo g2_holo \g1 z\0\ \g2 z\0\ apply (elim holomorphic_factor_unique) by (auto simp add:r_def) qed have P_exist:"\ n g r. P h n g r" when "\z'. (h \ z') (at z)" "isolated_singularity_at h z" "\\<^sub>Fw in (at z). h w\0" for h proof - from that(2) obtain r where "r>0" "h analytic_on ball z r - {z}" unfolding isolated_singularity_at_def by auto obtain z' where "(h \ z') (at z)" using \\z'. (h \ z') (at z)\ by auto define h' where "h'=(\x. if x=z then z' else h x)" have "h' holomorphic_on ball z r" apply (rule no_isolated_singularity'[of "{z}"]) subgoal by (metis LIM_equal Lim_at_imp_Lim_at_within \h \z\ z'\ empty_iff h'_def insert_iff) subgoal using \h analytic_on ball z r - {z}\ analytic_imp_holomorphic h'_def holomorphic_transform by fastforce by auto have ?thesis when "z'=0" proof - have "h' z=0" using that unfolding h'_def by auto moreover have "\ h' constant_on ball z r" using \\\<^sub>Fw in (at z). h w\0\ unfolding constant_on_def frequently_def eventually_at h'_def apply simp by (metis \0 < r\ centre_in_ball dist_commute mem_ball that) moreover note \h' holomorphic_on ball z r\ ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \ ball z r" and g:"g holomorphic_on ball z r1" "\w. w \ ball z r1 \ h' w = (w - z) ^ n * g w" "\w. w \ ball z r1 \ g w \ 0" using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified, OF \h' holomorphic_on ball z r\ \r>0\ \h' z=0\ \\ h' constant_on ball z r\] by (auto simp add:dist_commute) define rr where "rr=r1/2" have "P h' n g rr" unfolding P_def rr_def using \n>0\ \r1>0\ g by (auto simp add:powr_nat) then have "P h n g rr" unfolding h'_def P_def by auto then show ?thesis unfolding P_def by blast qed moreover have ?thesis when "z'\0" proof - have "h' z\0" using that unfolding h'_def by auto obtain r1 where "r1>0" "cball z r1 \ ball z r" "\x\cball z r1. h' x\0" proof - have "isCont h' z" "h' z\0" by (auto simp add: Lim_cong_within \h \z\ z'\ \z'\0\ continuous_at h'_def) then obtain r2 where r2:"r2>0" "\x\ball z r2. h' x\0" using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto define r1 where "r1=min r2 r / 2" have "0 < r1" "cball z r1 \ ball z r" using \r2>0\ \r>0\ unfolding r1_def by auto moreover have "\x\cball z r1. h' x \ 0" using r2 unfolding r1_def by simp ultimately show ?thesis using that by auto qed then have "P h' 0 h' r1" using \h' holomorphic_on ball z r\ unfolding P_def by auto then have "P h 0 h' r1" unfolding P_def h'_def by auto then show ?thesis unfolding P_def by blast qed ultimately show ?thesis by auto qed have ?thesis when "\x. (f \ x) (at z)" apply (rule_tac imp_unique[unfolded P_def]) using P_exist[OF that(1) f_iso non_zero] unfolding P_def . moreover have ?thesis when "is_pole f z" proof (rule imp_unique[unfolded P_def]) obtain e where [simp]:"e>0" and e_holo:"f holomorphic_on ball z e - {z}" and e_nz: "\x\ball z e-{z}. f x\0" proof - have "\\<^sub>F z in at z. f z \ 0" using \is_pole f z\ filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def by auto then obtain e1 where e1:"e1>0" "\x\ball z e1-{z}. f x\0" using that eventually_at[of "\x. f x\0" z UNIV,simplified] by (auto simp add:dist_commute) obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto define e where "e=min e1 e2" show ?thesis apply (rule that[of e]) using e1 e2 unfolding e_def by auto qed define h where "h \ \x. inverse (f x)" have "\n g r. P h n g r" proof - have "h \z\ 0" using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce moreover have "\\<^sub>Fw in (at z). h w\0" using non_zero apply (elim frequently_rev_mp) unfolding h_def eventually_at by (auto intro:exI[where x=1]) moreover have "isolated_singularity_at h z" unfolding isolated_singularity_at_def h_def apply (rule exI[where x=e]) using e_holo e_nz \e>0\ by (metis open_ball analytic_on_open holomorphic_on_inverse open_delete) ultimately show ?thesis using P_exist[of h] by auto qed then obtain n g r where "0 < r" and g_holo:"g holomorphic_on cball z r" and "g z\0" and g_fac:"(\w\cball z r-{z}. h w = g w * (w - z) powr of_int n \ g w \ 0)" unfolding P_def by auto have "P f (-n) (inverse o g) r" proof - have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\cball z r - {z}" for w using g_fac[rule_format,of w] that unfolding h_def apply (auto simp add:powr_minus ) by (metis inverse_inverse_eq inverse_mult_distrib) then show ?thesis unfolding P_def comp_def using \r>0\ g_holo g_fac \g z\0\ by (auto intro:holomorphic_intros) qed then show "\x g r. 0 < r \ g holomorphic_on cball z r \ g z \ 0 \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int x \ g w \ 0)" unfolding P_def by blast qed ultimately show ?thesis using \not_essential f z\ unfolding not_essential_def by presburger qed lemma not_essential_transform: assumes "not_essential g z" assumes "\\<^sub>F w in (at z). g w = f w" shows "not_essential f z" using assms unfolding not_essential_def by (simp add: filterlim_cong is_pole_cong) lemma isolated_singularity_at_transform: assumes "isolated_singularity_at g z" assumes "\\<^sub>F w in (at z). g w = f w" shows "isolated_singularity_at f z" proof - obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}" using assms(1) unfolding isolated_singularity_at_def by auto obtain r2 where "r2>0" and r2:" \x. x \ z \ dist x z < r2 \ g x = f x" using assms(2) unfolding eventually_at by auto define r3 where "r3=min r1 r2" have "r3>0" unfolding r3_def using \r1>0\ \r2>0\ by auto moreover have "f analytic_on ball z r3 - {z}" proof - have "g holomorphic_on ball z r3 - {z}" using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto) then have "f holomorphic_on ball z r3 - {z}" using r2 unfolding r3_def by (auto simp add:dist_commute elim!:holomorphic_transform) then show ?thesis by (subst analytic_on_open,auto) qed ultimately show ?thesis unfolding isolated_singularity_at_def by auto qed lemma not_essential_powr[singularity_intros]: assumes "LIM w (at z). f w :> (at x)" shows "not_essential (\w. (f w) powr (of_int n)) z" proof - define fp where "fp=(\w. (f w) powr (of_int n))" have ?thesis when "n>0" proof - have "(\w. (f w) ^ (nat n)) \z\ x ^ nat n" using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros) then have "fp \z\ x ^ nat n" unfolding fp_def apply (elim Lim_transform_within[where d=1],simp) by (metis less_le powr_0 powr_of_int that zero_less_nat_eq zero_power) then show ?thesis unfolding not_essential_def fp_def by auto qed moreover have ?thesis when "n=0" proof - have "fp \z\ 1 " apply (subst tendsto_cong[where g="\_.1"]) using that filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def by auto then show ?thesis unfolding fp_def not_essential_def by auto qed moreover have ?thesis when "n<0" proof (cases "x=0") case True have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity" apply (subst filterlim_inverse_at_iff[symmetric],simp) apply (rule filterlim_atI) subgoal using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros) subgoal using filterlim_at_within_not_equal[OF assms,of 0] by (eventually_elim,insert that,auto) done then have "LIM w (at z). fp w :> at_infinity" proof (elim filterlim_mono_eventually) show "\\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x" using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def apply eventually_elim using powr_of_int that by auto qed auto then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto next case False let ?xx= "inverse (x ^ (nat (-n)))" have "(\w. inverse ((f w) ^ (nat (-n)))) \z\?xx" using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros) then have "fp \z\?xx" apply (elim Lim_transform_within[where d=1],simp) unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less not_le power_eq_0_iff powr_0 powr_of_int that) then show ?thesis unfolding fp_def not_essential_def by auto qed ultimately show ?thesis by linarith qed lemma isolated_singularity_at_powr[singularity_intros]: assumes "isolated_singularity_at f z" "\\<^sub>F w in (at z). f w\0" shows "isolated_singularity_at (\w. (f w) powr (of_int n)) z" proof - obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}" using assms(1) unfolding isolated_singularity_at_def by auto then have r1:"f holomorphic_on ball z r1 - {z}" using analytic_on_open[of "ball z r1-{z}" f] by blast obtain r2 where "r2>0" and r2:"\w. w \ z \ dist w z < r2 \ f w \ 0" using assms(2) unfolding eventually_at by auto define r3 where "r3=min r1 r2" have "(\w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}" apply (rule holomorphic_on_powr_of_int) subgoal unfolding r3_def using r1 by auto subgoal unfolding r3_def using r2 by (auto simp add:dist_commute) done moreover have "r3>0" unfolding r3_def using \0 < r1\ \0 < r2\ by linarith ultimately show ?thesis unfolding isolated_singularity_at_def apply (subst (asm) analytic_on_open[symmetric]) by auto qed lemma non_zero_neighbour: assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows "\\<^sub>F w in (at z). f w\0" proof - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto have "f w \ 0" when " w \ z" "dist w z < fr" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w \ 0" using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute) moreover have "(w - z) powr of_int fn \0" unfolding powr_eq_0_iff using \w\z\ by auto ultimately show ?thesis by auto qed then show ?thesis using \fr>0\ unfolding eventually_at by auto qed lemma non_zero_neighbour_pole: assumes "is_pole f z" shows "\\<^sub>F w in (at z). f w\0" using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0] unfolding is_pole_def by auto lemma non_zero_neighbour_alt: assumes holo: "f holomorphic_on S" and "open S" "connected S" "z \ S" "\ \ S" "f \ \ 0" shows "\\<^sub>F w in (at z). f w\0 \ w\S" proof (cases "f z = 0") case True from isolated_zeros[OF holo \open S\ \connected S\ \z \ S\ True \\ \ S\ \f \ \ 0\] obtain r where "0 < r" "ball z r \ S" "\w \ ball z r - {z}.f w \ 0" by metis then show ?thesis unfolding eventually_at apply (rule_tac x=r in exI) by (auto simp add:dist_commute) next case False obtain r1 where r1:"r1>0" "\y. dist z y < r1 \ f y \ 0" using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at holo holomorphic_on_imp_continuous_on by blast obtain r2 where r2:"r2>0" "ball z r2 \ S" using assms(2) assms(4) openE by blast show ?thesis unfolding eventually_at apply (rule_tac x="min r1 r2" in exI) using r1 r2 by (auto simp add:dist_commute) qed lemma not_essential_times[singularity_intros]: assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "not_essential (\w. f w * g w) z" proof - define fg where "fg = (\w. f w * g w)" have ?thesis when "\ ((\\<^sub>Fw in (at z). f w\0) \ (\\<^sub>Fw in (at z). g w\0))" proof - have "\\<^sub>Fw in (at z). fg w=0" using that[unfolded frequently_def, simplified] unfolding fg_def by (auto elim: eventually_rev_mp) from tendsto_cong[OF this] have "fg \z\0" by auto then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" and g_nconst:"\\<^sub>Fw in (at z). g w\0" proof - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto obtain gn gp gr where [simp]:"gp z \ 0" and "gr > 0" and gr: "gp holomorphic_on cball z gr" "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto define r1 where "r1=(min fr gr)" have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" when "w\ball z r1 - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that unfolding r1_def by auto moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" using gr(2)[rule_format, of w] that unfolding r1_def by auto ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" unfolding fg_def by (auto simp add:powr_add) qed have [intro]: "fp \z\fp z" "gp \z\gp z" using fr(1) \fr>0\ gr(1) \gr>0\ by (meson open_ball ball_subset_cball centre_in_ball continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on holomorphic_on_subset)+ have ?thesis when "fn+gn>0" proof - have "(\w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \z\0" using that by (auto intro!:tendsto_eq_intros) then have "fg \z\ 0" apply (elim Lim_transform_within[OF _ \r1>0\]) by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int that) then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when "fn+gn=0" proof - have "(\w. fp w * gp w) \z\fp z*gp z" using that by (auto intro!:tendsto_eq_intros) then have "fg \z\ fp z*gp z" apply (elim Lim_transform_within[OF _ \r1>0\]) apply (subst fg_times) by (auto simp add:dist_commute that) then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when "fn+gn<0" proof - have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity" apply (rule filterlim_divide_at_infinity) apply (insert that, auto intro!:tendsto_eq_intros filterlim_atI) using eventually_at_topological by blast then have "is_pole fg z" unfolding is_pole_def apply (elim filterlim_transform_within[OF _ _ \r1>0\],simp) apply (subst fg_times,simp add:dist_commute) apply (subst powr_of_int) using that by (auto simp add:field_split_simps) then show ?thesis unfolding not_essential_def fg_def by auto qed ultimately show ?thesis unfolding not_essential_def fg_def by fastforce qed ultimately show ?thesis by auto qed lemma not_essential_inverse[singularity_intros]: assumes f_ness:"not_essential f z" assumes f_iso:"isolated_singularity_at f z" shows "not_essential (\w. inverse (f w)) z" proof - define vf where "vf = (\w. inverse (f w))" have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" proof - have "\\<^sub>Fw in (at z). f w=0" using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) then have "\\<^sub>Fw in (at z). vf w=0" unfolding vf_def by auto from tendsto_cong[OF this] have "vf \z\0" unfolding vf_def by auto then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "is_pole f z" proof - have "vf \z\0" using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "\x. f\z\x " and f_nconst:"\\<^sub>Fw in (at z). f w\0" proof - from that obtain fz where fz:"f\z\fz" by auto have ?thesis when "fz=0" proof - have "(\w. inverse (vf w)) \z\0" using fz that unfolding vf_def by auto moreover have "\\<^sub>F w in at z. inverse (vf w) \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] unfolding vf_def by auto ultimately have "is_pole vf z" using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "fz\0" proof - have "vf \z\inverse fz" using fz that unfolding vf_def by (auto intro:tendsto_eq_intros) then show ?thesis unfolding not_essential_def vf_def by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis using f_ness unfolding not_essential_def by auto qed lemma isolated_singularity_at_inverse[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" shows "isolated_singularity_at (\w. inverse (f w)) z" proof - define vf where "vf = (\w. inverse (f w))" have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" proof - have "\\<^sub>Fw in (at z). f w=0" using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) then have "\\<^sub>Fw in (at z). vf w=0" unfolding vf_def by auto then obtain d1 where "d1>0" and d1:"\x. x \ z \ dist x z < d1 \ vf x = 0" unfolding eventually_at by auto then have "vf holomorphic_on ball z d1-{z}" apply (rule_tac holomorphic_transform[of "\_. 0"]) by (auto simp add:dist_commute) then have "vf analytic_on ball z d1 - {z}" by (simp add: analytic_on_open open_delete) then show ?thesis using \d1>0\ unfolding isolated_singularity_at_def vf_def by auto qed moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" proof - have "\\<^sub>F w in at z. f w \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] . then obtain d1 where d1:"d1>0" "\x. x \ z \ dist x z < d1 \ f x \ 0" unfolding eventually_at by auto obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}" using f_iso unfolding isolated_singularity_at_def by auto define d3 where "d3=min d1 d2" have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto moreover have "vf analytic_on ball z d3 - {z}" unfolding vf_def apply (rule analytic_on_inverse) subgoal using d2 unfolding d3_def by (elim analytic_on_subset) auto subgoal for w using d1 unfolding d3_def by (auto simp add:dist_commute) done ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto qed ultimately show ?thesis by auto qed lemma not_essential_divide[singularity_intros]: assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "not_essential (\w. f w / g w) z" proof - have "not_essential (\w. f w * inverse (g w)) z" apply (rule not_essential_times[where g="\w. inverse (g w)"]) using assms by (auto intro: isolated_singularity_at_inverse not_essential_inverse) then show ?thesis by (simp add:field_simps) qed lemma assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows isolated_singularity_at_times[singularity_intros]: "isolated_singularity_at (\w. f w * g w) z" and isolated_singularity_at_add[singularity_intros]: "isolated_singularity_at (\w. f w + g w) z" proof - obtain d1 d2 where "d1>0" "d2>0" and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}" using f_iso g_iso unfolding isolated_singularity_at_def by auto define d3 where "d3=min d1 d2" have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto have "(\w. f w * g w) analytic_on ball z d3 - {z}" apply (rule analytic_on_mult) using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) then show "isolated_singularity_at (\w. f w * g w) z" using \d3>0\ unfolding isolated_singularity_at_def by auto have "(\w. f w + g w) analytic_on ball z d3 - {z}" apply (rule analytic_on_add) using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) then show "isolated_singularity_at (\w. f w + g w) z" using \d3>0\ unfolding isolated_singularity_at_def by auto qed lemma isolated_singularity_at_uminus[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" shows "isolated_singularity_at (\w. - f w) z" using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast lemma isolated_singularity_at_id[singularity_intros]: "isolated_singularity_at (\w. w) z" unfolding isolated_singularity_at_def by (simp add: gt_ex) lemma isolated_singularity_at_minus[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "isolated_singularity_at (\w. f w - g w) z" using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\w. - g w"] ,OF g_iso] by simp lemma isolated_singularity_at_divide[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and g_ness:"not_essential g z" shows "isolated_singularity_at (\w. f w / g w) z" using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso, of "\w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps) lemma isolated_singularity_at_const[singularity_intros]: "isolated_singularity_at (\w. c) z" unfolding isolated_singularity_at_def by (simp add: gt_ex) lemma isolated_singularity_at_holomorphic: assumes "f holomorphic_on s-{z}" "open s" "z\s" shows "isolated_singularity_at f z" using assms unfolding isolated_singularity_at_def by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) subsubsection \The order of non-essential singularities (i.e. removable singularities or poles)\ definition\<^marker>\tag important\ zorder :: "(complex \ complex) \ complex \ int" where "zorder f z = (THE n. (\h r. r>0 \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. f w = h w * (w-z) powr (of_int n) \ h w \0)))" definition\<^marker>\tag important\ zor_poly ::"[complex \ complex, complex] \ complex \ complex" where "zor_poly f z = (SOME h. \r. r > 0 \ h holomorphic_on cball z r \ h z \ 0 \ (\w\cball z r - {z}. f w = h w * (w - z) powr (zorder f z) \ h w \0))" lemma zorder_exist: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows "g z\0 \ (\r. r>0 \ g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w-z) powr n \ g w \0))" proof - define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have "\!n. \g r. P n g r" using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto then have "\g r. P n g r" unfolding n_def P_def zorder_def by (drule_tac theI',argo) then have "\r. P n g r" unfolding P_def zor_poly_def g_def n_def by (drule_tac someI_ex,argo) then obtain r1 where "P n g r1" by auto then show ?thesis unfolding P_def by auto qed lemma fixes f::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows zorder_inverse: "zorder (\w. inverse (f w)) z = - zorder f z" and zor_poly_inverse: "\\<^sub>Fw in (at z). zor_poly (\w. inverse (f w)) z w = inverse (zor_poly f z w)" proof - define vf where "vf = (\w. inverse (f w))" define fn vfn where "fn = zorder f z" and "vfn = zorder vf z" define fp vfp where "fp = zor_poly f z" and "vfp = zor_poly vf z" obtain fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def] by auto have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" and fr_nz: "inverse (fp w)\0" when "w\ball z fr - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that by auto then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\0" unfolding vf_def by (auto simp add:powr_minus) qed obtain vfr where [simp]:"vfp z \ 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr" "(\w\cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0)" proof - have "isolated_singularity_at vf z" using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def . moreover have "not_essential vf z" using not_essential_inverse[OF f_ness f_iso] unfolding vf_def . moreover have "\\<^sub>F w in at z. vf w \ 0" using f_nconst unfolding vf_def by (auto elim:frequently_elim1) ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto qed define r1 where "r1 = min fr vfr" have "r1>0" using \fr>0\ \vfr>0\ unfolding r1_def by simp show "vfn = - fn" apply (rule holomorphic_factor_unique[of r1 vfp z "\w. inverse (fp w)" vf]) subgoal using \r1>0\ by simp subgoal by simp subgoal by simp subgoal proof (rule ballI) fix w assume "w \ ball z r1 - {z}" then have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" unfolding r1_def by auto from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] show "vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0 \ vf w = inverse (fp w) * (w - z) powr of_int (- fn) \ inverse (fp w) \ 0" by auto qed subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros) subgoal using fr unfolding r1_def by (auto intro!:holomorphic_intros) done have "vfp w = inverse (fp w)" when "w\ball z r1-{z}" for w proof - have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" "w\z" using that unfolding r1_def by auto from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \vfn = - fn\ \w\z\ show ?thesis by auto qed then show "\\<^sub>Fw in (at z). vfp w = inverse (fp w)" unfolding eventually_at using \r1>0\ apply (rule_tac x=r1 in exI) by (auto simp add:dist_commute) qed lemma fixes f g::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and f_ness:"not_essential f z" and g_ness:"not_essential g z" and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" shows zorder_times:"zorder (\w. f w * g w) z = zorder f z + zorder g z" and zor_poly_times:"\\<^sub>Fw in (at z). zor_poly (\w. f w * g w) z w = zor_poly f z w *zor_poly g z w" proof - define fg where "fg = (\w. f w * g w)" define fn gn fgn where "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z" define fp gp fgp where "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z" have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" using fg_nconst by (auto elim!:frequently_elim1) obtain fr where [simp]:"fp z \ 0" and "fr > 0" and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto obtain gr where [simp]:"gp z \ 0" and "gr > 0" and gr: "gp holomorphic_on cball z gr" "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto define r1 where "r1=min fr gr" have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" when "w\ball z r1 - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that unfolding r1_def by auto moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" using gr(2)[rule_format, of w] that unfolding r1_def by auto ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" unfolding fg_def by (auto simp add:powr_add) qed obtain fgr where [simp]:"fgp z \ 0" and "fgr > 0" and fgr: "fgp holomorphic_on cball z fgr" "\w\cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0" proof - have "fgp z \ 0 \ (\r>0. fgp holomorphic_on cball z r \ (\w\cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0))" apply (rule zorder_exist[of fg z, folded fgn_def fgp_def]) subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] . subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] . subgoal unfolding fg_def using fg_nconst . done then show ?thesis using that by blast qed define r2 where "r2 = min fgr r1" have "r2>0" using \r1>0\ \fgr>0\ unfolding r2_def by simp show "fgn = fn + gn " apply (rule holomorphic_factor_unique[of r2 fgp z "\w. fp w * gp w" fg]) subgoal using \r2>0\ by simp subgoal by simp subgoal by simp subgoal proof (rule ballI) fix w assume "w \ ball z r2 - {z}" then have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" unfolding r2_def by auto from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] show "fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0 \ fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \ fp w * gp w \ 0" by auto qed subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) done have "fgp w = fp w *gp w" when "w\ball z r2-{z}" for w proof - have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" "w\z" using that unfolding r2_def by auto from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \fgn = fn + gn\ \w\z\ show ?thesis by auto qed then show "\\<^sub>Fw in (at z). fgp w = fp w * gp w" using \r2>0\ unfolding eventually_at by (auto simp add:dist_commute) qed lemma fixes f g::"complex \ complex" and z::complex assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and f_ness:"not_essential f z" and g_ness:"not_essential g z" and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" shows zorder_divide:"zorder (\w. f w / g w) z = zorder f z - zorder g z" and zor_poly_divide:"\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" proof - have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" using fg_nconst by (auto elim!:frequently_elim1) define vg where "vg=(\w. inverse (g w))" have "zorder (\w. f w * vg w) z = zorder f z + zorder vg z" apply (rule zorder_times[OF f_iso _ f_ness,of vg]) subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) done then show "zorder (\w. f w / g w) z = zorder f z - zorder g z" using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def by (auto simp add:field_simps) have "\\<^sub>F w in at z. zor_poly (\w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w" apply (rule zor_poly_times[OF f_iso _ f_ness,of vg]) subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) done then show "\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def apply eventually_elim by (auto simp add:field_simps) qed lemma zorder_exist_zero: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes holo: "f holomorphic_on s" and "open s" "connected s" "z\s" and non_const: "\w\s. f w \ 0" shows "(if f z=0 then n > 0 else n=0) \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r \ (\w\cball z r. f w = g w * (w-z) ^ nat n \ g w \0))" proof - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" proof - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" proof (rule zorder_exist[of f z,folded g_def n_def]) show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using holo assms(4,6) by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE) show "not_essential f z" unfolding not_essential_def using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce have "\\<^sub>F w in at z. f w \ 0 \ w\s" proof - obtain w where "w\s" "f w\0" using non_const by auto then show ?thesis by (rule non_zero_neighbour_alt[OF holo \open s\ \connected s\ \z\s\]) qed then show "\\<^sub>F w in at z. f w \ 0" apply (elim eventually_frequentlyE) by auto qed then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" by auto obtain r2 where r2: "r2>0" "cball z r2 \ s" using assms(4,6) open_contains_cball_eq by blast define r3 where "r3=min r1 r2" have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto moreover have "g holomorphic_on cball z r3" using r1(1) unfolding r3_def by auto moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" using r1(2) unfolding r3_def by auto ultimately show ?thesis using that[of r3] \g z\0\ by auto qed have if_0:"if f z=0 then n > 0 else n=0" proof - have "f\ z \ f z" by (metis assms(4,6,7) at_within_open continuous_on holo holomorphic_on_imp_continuous_on) then have "(\w. g w * (w - z) powr of_int n) \z\ f z" apply (elim Lim_transform_within_open[where s="ball z r"]) using r by auto moreover have "g \z\g z" by (metis (mono_tags, lifting) open_ball at_within_open_subset ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1,3) subsetCE) ultimately have "(\w. (g w * (w - z) powr of_int n) / g w) \z\ f z/g z" apply (rule_tac tendsto_divide) using \g z\0\ by auto then have powr_tendsto:"(\w. (w - z) powr of_int n) \z\ f z/g z" apply (elim Lim_transform_within_open[where s="ball z r"]) using r by auto have ?thesis when "n\0" "f z=0" proof - have "(\w. (w - z) ^ nat n) \z\ f z/g z" using powr_tendsto apply (elim Lim_transform_within[where d=r]) by (auto simp add: powr_of_int \n\0\ \r>0\) then have *:"(\w. (w - z) ^ nat n) \z\ 0" using \f z=0\ by simp moreover have False when "n=0" proof - have "(\w. (w - z) ^ nat n) \z\ 1" using \n=0\ by auto then show False using * using LIM_unique zero_neq_one by blast qed ultimately show ?thesis using that by fastforce qed moreover have ?thesis when "n\0" "f z\0" proof - have False when "n>0" proof - have "(\w. (w - z) ^ nat n) \z\ f z/g z" using powr_tendsto apply (elim Lim_transform_within[where d=r]) by (auto simp add: powr_of_int \n\0\ \r>0\) moreover have "(\w. (w - z) ^ nat n) \z\ 0" using \n>0\ by (auto intro!:tendsto_eq_intros) ultimately show False using \f z\0\ \g z\0\ using LIM_unique divide_eq_0_iff by blast qed then show ?thesis using that by force qed moreover have False when "n<0" proof - have "(\w. inverse ((w - z) ^ nat (- n))) \z\ f z/g z" "(\w.((w - z) ^ nat (- n))) \z\ 0" subgoal using powr_tendsto powr_of_int that by (elim Lim_transform_within_open[where s=UNIV],auto) subgoal using that by (auto intro!:tendsto_eq_intros) done from tendsto_mult[OF this,simplified] have "(\x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \z\ 0" . then have "(\x. 1::complex) \z\ 0" by (elim Lim_transform_within_open[where s=UNIV],auto) then show False using LIM_const_eq by fastforce qed ultimately show ?thesis by fastforce qed moreover have "f w = g w * (w-z) ^ nat n \ g w \0" when "w\cball z r" for w proof (cases "w=z") case True then have "f \z\f w" using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce then have "(\w. g w * (w-z) ^ nat n) \z\f w" proof (elim Lim_transform_within[OF _ \r>0\]) fix x assume "0 < dist x z" "dist x z < r" then have "x \ cball z r - {z}" "x\z" unfolding cball_def by (auto simp add: dist_commute) then have "f x = g x * (x - z) powr of_int n" using r(4)[rule_format,of x] by simp also have "... = g x * (x - z) ^ nat n" apply (subst powr_of_int) using if_0 \x\z\ by (auto split:if_splits) finally show "f x = g x * (x - z) ^ nat n" . qed moreover have "(\w. g w * (w-z) ^ nat n) \z\ g w * (w-z) ^ nat n" using True apply (auto intro!:tendsto_eq_intros) by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1) r(3) that) ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast then show ?thesis using \g z\0\ True by auto next case False then have "f w = g w * (w - z) powr of_int n \ g w \ 0" using r(4) that by auto then show ?thesis using False if_0 powr_of_int by (auto split:if_splits) qed ultimately show ?thesis using r by auto qed lemma zorder_exist_pole: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" assumes holo: "f holomorphic_on s-{z}" and "open s" "z\s" and "is_pole f z" shows "n < 0 \ g z\0 \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0))" proof - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" proof - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" proof (rule zorder_exist[of f z,folded g_def n_def]) show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using holo assms(4,5) by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) show "not_essential f z" unfolding not_essential_def using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce from non_zero_neighbour_pole[OF \is_pole f z\] show "\\<^sub>F w in at z. f w \ 0" apply (elim eventually_frequentlyE) by auto qed then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" by auto obtain r2 where r2: "r2>0" "cball z r2 \ s" using assms(4,5) open_contains_cball_eq by metis define r3 where "r3=min r1 r2" have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto moreover have "g holomorphic_on cball z r3" using r1(1) unfolding r3_def by auto moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" using r1(2) unfolding r3_def by auto ultimately show ?thesis using that[of r3] \g z\0\ by auto qed have "n<0" proof (rule ccontr) assume " \ n < 0" define c where "c=(if n=0 then g z else 0)" have [simp]:"g \z\ g z" by (metis open_ball at_within_open ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) ) have "\\<^sub>F x in at z. f x = g x * (x - z) ^ nat n" unfolding eventually_at_topological apply (rule_tac exI[where x="ball z r"]) using r powr_of_int \\ n < 0\ by auto moreover have "(\x. g x * (x - z) ^ nat n) \z\c" proof (cases "n=0") case True then show ?thesis unfolding c_def by simp next case False then have "(\x. (x - z) ^ nat n) \z\ 0" using \\ n < 0\ by (auto intro!:tendsto_eq_intros) from tendsto_mult[OF _ this,of g "g z",simplified] show ?thesis unfolding c_def using False by simp qed ultimately have "f \z\c" using tendsto_cong by fast then show False using \is_pole f z\ at_neq_bot not_tendsto_and_filterlim_at_infinity unfolding is_pole_def by blast qed moreover have "\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0" using r(4) \n<0\ powr_of_int by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le) ultimately show ?thesis using r(1-3) \g z\0\ by auto qed lemma zorder_eqI: assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" assumes fg_eq:"\w. \w \ s;w\z\ \ f w = g w * (w - z) powr n" shows "zorder f z = n" proof - have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact moreover have "open (-{0::complex})" by auto ultimately have "open ((g -` (-{0})) \ s)" unfolding continuous_on_open_vimage[OF \open s\] by blast moreover from assms have "z \ (g -` (-{0})) \ s" by auto ultimately obtain r where r: "r > 0" "cball z r \ s \ (g -` (-{0}))" unfolding open_contains_cball by blast let ?gg= "(\w. g w * (w - z) powr n)" define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have "P n g r" unfolding P_def using r assms(3,4,5) by auto then have "\g r. P n g r" by auto moreover have unique: "\!n. \g r. P n g r" unfolding P_def proof (rule holomorphic_factor_puncture) have "ball z r-{z} \ s" using r using ball_subset_cball by blast then have "?gg holomorphic_on ball z r-{z}" using \g holomorphic_on s\ r by (auto intro!: holomorphic_intros) then have "f holomorphic_on ball z r - {z}" apply (elim holomorphic_transform) using fg_eq \ball z r-{z} \ s\ by auto then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using analytic_on_open open_delete r(1) by blast next have "not_essential ?gg z" proof (intro singularity_intros) show "not_essential g z" by (meson \continuous_on s g\ assms(1) assms(2) continuous_on_eq_continuous_at isCont_def not_essential_def) show " \\<^sub>F w in at z. w - z \ 0" by (simp add: eventually_at_filter) then show "LIM w at z. w - z :> at 0" unfolding filterlim_at by (auto intro:tendsto_eq_intros) show "isolated_singularity_at g z" by (meson Diff_subset open_ball analytic_on_holomorphic assms(1,2,3) holomorphic_on_subset isolated_singularity_at_def openE) qed then show "not_essential f z" apply (elim not_essential_transform) unfolding eventually_at using assms(1,2) assms(5)[symmetric] by (metis dist_commute mem_ball openE subsetCE) show "\\<^sub>F w in at z. f w \ 0" unfolding frequently_at proof (rule,rule) fix d::real assume "0 < d" define z' where "z'=z+min d r / 2" have "z' \ z" " dist z' z < d " unfolding z'_def using \d>0\ \r>0\ by (auto simp add:dist_norm) moreover have "f z' \ 0" proof (subst fg_eq[OF _ \z'\z\]) have "z' \ cball z r" unfolding z'_def using \r>0\ \d>0\ by (auto simp add:dist_norm) then show " z' \ s" using r(2) by blast show "g z' * (z' - z) powr of_int n \ 0" using P_def \P n g r\ \z' \ cball z r\ calculation(1) by auto qed ultimately show "\x\UNIV. x \ z \ dist x z < d \ f x \ 0" by auto qed qed ultimately have "(THE n. \g r. P n g r) = n" by (rule_tac the1_equality) then show ?thesis unfolding zorder_def P_def by blast qed lemma residue_pole_order: fixes f::"complex \ complex" and z::complex defines "n \ nat (- zorder f z)" and "h \ zor_poly f z" assumes f_iso:"isolated_singularity_at f z" and pole:"is_pole f z" shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ ball z e" and h_holo: "h holomorphic_on cball z r" and h_divide:"(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" proof - obtain r where r:"zorder f z < 0" "h z \ 0" "r>0" "cball z r \ ball z e" "h holomorphic_on cball z r" "(\w\cball z r - {z}. f w = h w / (w - z) ^ n \ h w \ 0)" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\,folded n_def h_def] by auto have "n>0" using \zorder f z < 0\ unfolding n_def by simp moreover have "(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" using \h z\0\ r(6) by blast ultimately show ?thesis using r(3,4,5) that by blast qed have r_nonzero:"\w. w \ ball z r - {z} \ f w \ 0" using h_divide by simp define c where "c \ 2 * pi * \" define der_f where "der_f \ ((deriv ^^ (n - 1)) h z / fact (n-1))" define h' where "h' \ \u. h u / (u - z) ^ n" have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" unfolding h'_def proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", folded c_def Suc_pred'[OF \n>0\]]) show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp show "h holomorphic_on ball z r" using h_holo by auto show " z \ ball z r" using \r>0\ by auto qed then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto then have "(f has_contour_integral c * der_f) (circlepath z r)" proof (elim has_contour_integral_eq) fix x assume "x \ path_image (circlepath z r)" hence "x\cball z r - {z}" using \r>0\ by auto then show "h' x = f x" using h_divide unfolding h'_def by auto qed moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" using base_residue[of \ball z e\ z,simplified,OF \r>0\ f_holo r_cball,folded c_def] unfolding c_def by simp ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast hence "der_f = residue f z" unfolding c_def by auto thus ?thesis unfolding der_f_def by auto qed lemma simple_zeroI: assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" assumes "\w. w \ s \ f w = g w * (w - z)" shows "zorder f z = 1" using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto) lemma higher_deriv_power: shows "(deriv ^^ j) (\w. (w - z) ^ n) w = pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)" proof (induction j arbitrary: w) case 0 thus ?case by auto next case (Suc j w) have "(deriv ^^ Suc j) (\w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\w. (w - z) ^ n)) w" by simp also have "(deriv ^^ j) (\w. (w - z) ^ n) = (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))" using Suc by (intro Suc.IH ext) also { have "(\ has_field_derivative of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)" using Suc.prems by (auto intro!: derivative_eq_intros) also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = pochhammer (of_nat (Suc n - Suc j)) (Suc j)" by (cases "Suc j \ n", subst pochhammer_rec) (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left) finally have "deriv (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w = \ * (w - z) ^ (n - Suc j)" by (rule DERIV_imp_deriv) } finally show ?case . qed lemma zorder_zero_eqI: assumes f_holo:"f holomorphic_on s" and "open s" "z \ s" assumes zero: "\i. i < nat n \ (deriv ^^ i) f z = 0" assumes nz: "(deriv ^^ nat n) f z \ 0" and "n\0" shows "zorder f z = n" proof - obtain r where [simp]:"r>0" and "ball z r \ s" using \open s\ \z\s\ openE by blast have nz':"\w\ball z r. f w \ 0" proof (rule ccontr) assume "\ (\w\ball z r. f w \ 0)" then have "eventually (\u. f u = 0) (nhds z)" using \r>0\ unfolding eventually_nhds apply (rule_tac x="ball z r" in exI) by auto then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\_. 0) z" by (intro higher_deriv_cong_ev) auto also have "(deriv ^^ nat n) (\_. 0) z = 0" by (induction n) simp_all finally show False using nz by contradiction qed define zn g where "zn = zorder f z" and "g = zor_poly f z" obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and [simp]:"e>0" and "cball z e \ ball z r" and g_holo:"g holomorphic_on cball z e" and e_fac:"(\w\cball z e. f w = g w * (w - z) ^ nat zn \ g w \ 0)" proof - have "f holomorphic_on ball z r" using f_holo \ball z r \ s\ by auto from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def] show ?thesis by blast qed from this(1,2,5) have "zn\0" "g z\0" subgoal by (auto split:if_splits) subgoal using \0 < e\ ball_subset_cball centre_in_ball e_fac by blast done define A where "A = (\i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)" have deriv_A:"(deriv ^^ i) f z = (if zn \ int i then A i else 0)" for i proof - have "eventually (\w. w \ ball z e) (nhds z)" using \cball z e \ ball z r\ \e>0\ by (intro eventually_nhds_in_open) auto hence "eventually (\w. f w = (w - z) ^ (nat zn) * g w) (nhds z)" apply eventually_elim by (use e_fac in auto) hence "(deriv ^^ i) f z = (deriv ^^ i) (\w. (w - z) ^ nat zn * g w) z" by (intro higher_deriv_cong_ev) auto also have "\ = (\j=0..i. of_nat (i choose j) * (deriv ^^ j) (\w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)" using g_holo \e>0\ by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros) also have "\ = (\j=0..i. if j = nat zn then of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" proof (intro sum.cong refl, goal_cases) case (1 j) have "(deriv ^^ j) (\w. (w - z) ^ nat zn) z = pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)" by (subst higher_deriv_power) auto also have "\ = (if j = nat zn then fact j else 0)" by (auto simp: not_less pochhammer_0_left pochhammer_fact) also have "of_nat (i choose j) * \ * (deriv ^^ (i - j)) g z = (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" by simp finally show ?case . qed also have "\ = (if i \ zn then A i else 0)" by (auto simp: A_def) finally show "(deriv ^^ i) f z = \" . qed have False when "nn\0\ by auto with nz show False by auto qed moreover have "n\zn" proof - have "g z \ 0" using e_fac[rule_format,of z] \e>0\ by simp then have "(deriv ^^ nat zn) f z \ 0" using deriv_A[of "nat zn"] by(auto simp add:A_def) then have "nat zn \ nat n" using zero[of "nat zn"] by linarith moreover have "zn\0" using e_if by (auto split:if_splits) ultimately show ?thesis using nat_le_eq_zle by blast qed ultimately show ?thesis unfolding zn_def by fastforce qed lemma assumes "eventually (\z. f z = g z) (at z)" "z = z'" shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'" proof - define P where "P = (\ff n h r. 0 < r \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \ h w\0))" have "(\r. P f n h r) = (\r. P g n h r)" for n h proof - have *: "\r. P g n h r" if "\r. P f n h r" and "eventually (\x. f x = g x) (at z)" for f g proof - from that(1) obtain r1 where r1_P:"P f n h r1" by auto from that(2) obtain r2 where "r2>0" and r2_dist:"\x. x \ z \ dist x z \ r2 \ f x = g x" unfolding eventually_at_le by auto define r where "r=min r1 r2" have "r>0" "h z\0" using r1_P \r2>0\ unfolding r_def P_def by auto moreover have "h holomorphic_on cball z r" using r1_P unfolding P_def r_def by auto moreover have "g w = h w * (w - z) powr of_int n \ h w \ 0" when "w\cball z r - {z}" for w proof - have "f w = h w * (w - z) powr of_int n \ h w \ 0" using r1_P that unfolding P_def r_def by auto moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def by (simp add: dist_commute) ultimately show ?thesis by simp qed ultimately show ?thesis unfolding P_def by auto qed from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) show ?thesis by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']]) qed then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'" using \z=z'\ unfolding P_def zorder_def zor_poly_def by auto qed lemma zorder_nonzero_div_power: assumes "open s" "z \ s" "f holomorphic_on s" "f z \ 0" "n > 0" shows "zorder (\w. f w / (w - z) ^ n) z = - n" apply (rule zorder_eqI[OF \open s\ \z\s\ \f holomorphic_on s\ \f z\0\]) apply (subst powr_of_int) using \n>0\ by (auto simp add:field_simps) lemma zor_poly_eq: assumes "isolated_singularity_at f z" "not_essential f z" "\\<^sub>F w in at z. f w \ 0" shows "eventually (\w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)" proof - obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))" using zorder_exist[OF assms] by blast then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z" by (auto simp: field_simps powr_minus) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_zero_eq: assumes "f holomorphic_on s" "open s" "connected s" "z \ s" "\w\s. f w \ 0" shows "eventually (\w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)" proof - obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))" using zorder_exist_zero[OF assms] by auto then have *: "\w\ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)" by (auto simp: field_simps powr_minus) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_pole_eq: assumes f_iso:"isolated_singularity_at f z" "is_pole f z" shows "eventually (\w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)" proof - obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\] by auto then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)" by (auto simp: field_simps) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes "isolated_singularity_at f z0" "not_essential f z0" "\\<^sub>F w in at z0. f w \ 0" assumes lim: "((\x. f (g x) * (g x - z0) powr - n) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - from zorder_exist[OF assms(2-4)] obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" "\w. w \ cball z0 r - {z0} \ f w = zor_poly f z0 w * (w - z0) powr n" unfolding n_def by blast from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto hence "eventually (\w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)" by eventually_elim (insert r, auto simp: field_simps powr_minus) moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w * (w - z0) powr - n) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) * (g x - z0) powr - n) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma zor_poly_zero_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes "f holomorphic_on A" "open A" "connected A" "z0 \ A" "\z\A. f z \ 0" assumes lim: "((\x. f (g x) / (g x - z0) ^ nat n) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - from zorder_exist_zero[OF assms(2-6)] obtain r where r: "r > 0" "cball z0 r \ A" "zor_poly f z0 holomorphic_on cball z0 r" "\w. w \ cball z0 r \ f w = zor_poly f z0 w * (w - z0) ^ nat n" unfolding n_def by blast from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto hence "eventually (\w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)" by eventually_elim (insert r, auto simp: field_simps) moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w / (w - z0) ^ nat n) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) / (g x - z0) ^ nat n) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma zor_poly_pole_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0" assumes lim: "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" proof - have "\\<^sub>F w in at z0. f w \ 0" using non_zero_neighbour_pole[OF \is_pole f z0\] by (auto elim:eventually_frequentlyE) moreover have "not_essential f z0" unfolding not_essential_def using \is_pole f z0\ by simp ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto qed from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto have "eventually (\w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)" using zor_poly_pole_eq[OF f_iso \is_pole f z0\] unfolding n_def . moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w * (w - z0) ^ nat (-n)) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma residue_simple_pole: assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" shows "residue f z0 = zor_poly f z0 z0" using assms by (subst residue_pole_order) simp_all lemma residue_simple_pole_limit: assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" assumes "((\x. f (g x) * (g x - z0)) \ c) F" assumes "filterlim g (at z0) F" "F \ bot" shows "residue f z0 = c" proof - have "residue f z0 = zor_poly f z0 z0" by (rule residue_simple_pole assms)+ also have "\ = c" apply (rule zor_poly_pole_eqI) using assms by auto finally show ?thesis . qed lemma lhopital_complex_simple: assumes "(f has_field_derivative f') (at z)" assumes "(g has_field_derivative g') (at z)" assumes "f z = 0" "g z = 0" "g' \ 0" "f' / g' = c" shows "((\w. f w / g w) \ c) (at z)" proof - have "eventually (\w. w \ z) (at z)" by (auto simp: eventually_at_filter) hence "eventually (\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)" by eventually_elim (simp add: assms field_split_simps) moreover have "((\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \ f' / g') (at z)" by (intro tendsto_divide has_field_derivativeD assms) ultimately have "((\w. f w / g w) \ f' / g') (at z)" by (blast intro: Lim_transform_eventually) with assms show ?thesis by simp qed lemma assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" and "open s" "connected s" "z \ s" assumes g_deriv:"(g has_field_derivative g') (at z)" assumes "f z \ 0" "g z = 0" "g' \ 0" shows porder_simple_pole_deriv: "zorder (\w. f w / g w) z = - 1" and residue_simple_pole_deriv: "residue (\w. f w / g w) z = f z / g'" proof - have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z" using isolated_singularity_at_holomorphic[OF _ \open s\ \z\s\] f_holo g_holo by (meson Diff_subset holomorphic_on_subset)+ have [simp]:"not_essential f z" "not_essential g z" unfolding not_essential_def using f_holo g_holo assms(3,5) by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+ have g_nconst:"\\<^sub>F w in at z. g w \0 " proof (rule ccontr) assume "\ (\\<^sub>F w in at z. g w \ 0)" then have "\\<^sub>F w in nhds z. g w = 0" unfolding eventually_at eventually_nhds frequently_at using \g z = 0\ by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball) then have "deriv g z = deriv (\_. 0) z" by (intro deriv_cong_ev) auto then have "deriv g z = 0" by auto then have "g' = 0" using g_deriv DERIV_imp_deriv by blast then show False using \g'\0\ by auto qed have "zorder (\w. f w / g w) z = zorder f z - zorder g z" proof - have "\\<^sub>F w in at z. f w \0 \ w\s" apply (rule non_zero_neighbour_alt) using assms by auto with g_nconst have "\\<^sub>F w in at z. f w * g w \ 0" by (elim frequently_rev_mp eventually_rev_mp,auto) then show ?thesis using zorder_divide[of f z g] by auto qed moreover have "zorder f z=0" apply (rule zorder_zero_eqI[OF f_holo \open s\ \z\s\]) using \f z\0\ by auto moreover have "zorder g z=1" apply (rule zorder_zero_eqI[OF g_holo \open s\ \z\s\]) subgoal using assms(8) by auto subgoal using DERIV_imp_deriv assms(9) g_deriv by auto subgoal by simp done ultimately show "zorder (\w. f w / g w) z = - 1" by auto show "residue (\w. f w / g w) z = f z / g'" proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified]) show "zorder (\w. f w / g w) z = - 1" by fact show "isolated_singularity_at (\w. f w / g w) z" by (auto intro: singularity_intros) show "is_pole (\w. f w / g w) z" proof (rule is_pole_divide) have "\\<^sub>F x in at z. g x \ 0" apply (rule non_zero_neighbour) using g_nconst by auto moreover have "g \z\ 0" using DERIV_isCont assms(8) continuous_at g_deriv by force ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp show "isCont f z" using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on by auto show "f z \ 0" by fact qed show "filterlim id (at z) (at z)" by (simp add: filterlim_iff) have "((\w. (f w * (w - z)) / g w) \ f z / g') (at z)" proof (rule lhopital_complex_simple) show "((\w. f w * (w - z)) has_field_derivative f z) (at z)" using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo]) show "(g has_field_derivative g') (at z)" by fact qed (insert assms, auto) then show "((\w. (f w / g w) * (w - z)) \ f z / g') (at z)" by (simp add: field_split_simps) qed qed subsection \The argument principle\ theorem argument_principle: fixes f::"complex \ complex" and poles s:: "complex set" defines "pz \ {w. f w = 0 \ w \ poles}" \ \\<^term>\pz\ is the set of poles and zeros\ assumes "open s" and "connected s" and f_holo:"f holomorphic_on s-poles" and h_holo:"h holomorphic_on s" and "valid_path g" and loop:"pathfinish g = pathstart g" and path_img:"path_image g \ s - pz" and homo:"\z. (z \ s) \ winding_number g z = 0" and finite:"finite pz" and poles:"\p\poles. is_pole f p" shows "contour_integral g (\x. deriv f x * h x / f x) = 2 * pi * \ * (\p\pz. winding_number g p * h p * zorder f p)" (is "?L=?R") proof - define c where "c \ 2 * complex_of_real pi * \ " define ff where "ff \ (\x. deriv f x * h x / f x)" define cont where "cont \ \ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" define avoid where "avoid \ \p e. \w\cball p e. w \ s \ (w \ p \ w \ pz)" have "\e>0. avoid p e \ (p\pz \ cont ff p e)" when "p\s" for p proof - obtain e1 where "e1>0" and e1_avoid:"avoid p e1" using finite_cball_avoid[OF \open s\ finite] \p\s\ unfolding avoid_def by auto have "\e2>0. cball p e2 \ ball p e1 \ cont ff p e2" when "p\pz" proof - define po where "po \ zorder f p" define pp where "pp \ zor_poly f p" define f' where "f' \ \w. pp w * (w - p) powr po" define ff' where "ff' \ (\x. deriv f' x * h x / f' x)" obtain r where "pp p\0" "r>0" and "rw\cball p r-{p}. f w = pp w * (w - p) powr po \ pp w \ 0)" proof - have "isolated_singularity_at f p" proof - have "f holomorphic_on ball p e1 - {p}" apply (intro holomorphic_on_subset[OF f_holo]) using e1_avoid \p\pz\ unfolding avoid_def pz_def by force then show ?thesis unfolding isolated_singularity_at_def using \e1>0\ analytic_on_open open_delete by blast qed moreover have "not_essential f p" proof (cases "is_pole f p") case True then show ?thesis unfolding not_essential_def by auto next case False then have "p\s-poles" using \p\s\ poles unfolding pz_def by auto moreover have "open (s-poles)" using \open s\ apply (elim open_Diff) apply (rule finite_imp_closed) using finite unfolding pz_def by simp ultimately have "isCont f p" using holomorphic_on_imp_continuous_on[OF f_holo] continuous_on_eq_continuous_at by auto then show ?thesis unfolding isCont_def not_essential_def by auto qed moreover have "\\<^sub>F w in at p. f w \ 0 " proof (rule ccontr) assume "\ (\\<^sub>F w in at p. f w \ 0)" then have "\\<^sub>F w in at p. f w= 0" unfolding frequently_def by auto then obtain rr where "rr>0" "\w\ball p rr - {p}. f w =0" unfolding eventually_at by (auto simp add:dist_commute) then have "ball p rr - {p} \ {w\ball p rr-{p}. f w=0}" by blast moreover have "infinite (ball p rr - {p})" using \rr>0\ using finite_imp_not_open by fastforce ultimately have "infinite {w\ball p rr-{p}. f w=0}" using infinite_super by blast then have "infinite pz" unfolding pz_def infinite_super by auto then show False using \finite pz\ by auto qed ultimately obtain r where "pp p \ 0" and r:"r>0" "pp holomorphic_on cball p r" "(\w\cball p r - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" using zorder_exist[of f p,folded po_def pp_def] by auto define r1 where "r1=min r e1 / 2" have "r1e1>0\ \r>0\ by auto moreover have "r1>0" "pp holomorphic_on cball p r1" "(\w\cball p r1 - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" unfolding r1_def using \e1>0\ r by auto ultimately show ?thesis using that \pp p\0\ by auto qed define e2 where "e2 \ r/2" have "e2>0" using \r>0\ unfolding e2_def by auto define anal where "anal \ \w. deriv pp w * h w / pp w" define prin where "prin \ \w. po * h w / (w - p)" have "((\w. prin w + anal w) has_contour_integral c * po * h p) (circlepath p e2)" proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) have "ball p r \ s" using \r avoid_def ball_subset_cball e1_avoid by (simp add: subset_eq) then have "cball p e2 \ s" using \r>0\ unfolding e2_def by auto then have "(\w. po * h w) holomorphic_on cball p e2" using h_holo by (auto intro!: holomorphic_intros) then show "(prin has_contour_integral c * po * h p ) (circlepath p e2)" using Cauchy_integral_circlepath_simple[folded c_def, of "\w. po * h w"] \e2>0\ unfolding prin_def by (auto simp add: mult.assoc) have "anal holomorphic_on ball p r" unfolding anal_def using pp_holo h_holo pp_po \ball p r \ s\ \pp p\0\ by (auto intro!: holomorphic_intros) then show "(anal has_contour_integral 0) (circlepath p e2)" using e2_def \r>0\ by (auto elim!: Cauchy_theorem_disc_simple) qed then have "cont ff' p e2" unfolding cont_def po_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto define wp where "wp \ w-p" have "wp\0" and "pp w \0" unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto moreover have der_f':"deriv f' w = po * pp w * (w-p) powr (po - 1) + deriv pp w * (w-p) powr po" proof (rule DERIV_imp_deriv) have "(pp has_field_derivative (deriv pp w)) (at w)" using DERIV_deriv_iff_has_field_derivative pp_holo \w\p\ by (meson open_ball \w \ ball p r\ ball_subset_cball holomorphic_derivI holomorphic_on_subset) then show " (f' has_field_derivative of_int po * pp w * (w - p) powr of_int (po - 1) + deriv pp w * (w - p) powr of_int po) (at w)" unfolding f'_def using \w\p\ by (auto intro!: derivative_eq_intros DERIV_cong[OF has_field_derivative_powr_of_int]) qed ultimately show "prin w + anal w = ff' w" unfolding ff'_def prin_def anal_def apply simp apply (unfold f'_def) apply (fold wp_def) apply (auto simp add:field_simps) by (metis (no_types, lifting) diff_add_cancel mult.commute powr_add powr_to_1) qed then have "cont ff p e2" unfolding cont_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto have "deriv f' w = deriv f w" proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo by (auto intro!: holomorphic_intros) next have "ball p e1 - {p} \ s - poles" using ball_subset_cball e1_avoid[unfolded avoid_def] unfolding pz_def by auto then have "ball p r - {p} \ s - poles" apply (elim dual_order.trans) using \r by auto then show "f holomorphic_on ball p r - {p}" using f_holo by auto next show "open (ball p r - {p})" by auto show "w \ ball p r - {p}" using \w\ball p r\ \w\p\ by auto next fix x assume "x \ ball p r - {p}" then show "f' x = f x" using pp_po unfolding f'_def by auto qed moreover have " f' w = f w " using \w \ ball p r\ ball_subset_cball subset_iff pp_po \w\p\ unfolding f'_def by auto ultimately show "ff' w = ff w" unfolding ff'_def ff_def by simp qed moreover have "cball p e2 \ ball p e1" using \0 < r\ \r e2_def by auto ultimately show ?thesis using \e2>0\ by auto qed then obtain e2 where e2:"p\pz \ e2>0 \ cball p e2 \ ball p e1 \ cont ff p e2" by auto define e4 where "e4 \ if p\pz then e2 else e1" have "e4>0" using e2 \e1>0\ unfolding e4_def by auto moreover have "avoid p e4" using e2 \e1>0\ e1_avoid unfolding e4_def avoid_def by auto moreover have "p\pz \ cont ff p e4" by (auto simp add: e2 e4_def) ultimately show ?thesis by auto qed then obtain get_e where get_e:"\p\s. get_e p>0 \ avoid p (get_e p) \ (p\pz \ cont ff p (get_e p))" by metis define ci where "ci \ \p. contour_integral (circlepath p (get_e p)) ff" define w where "w \ \p. winding_number g p" have "contour_integral g ff = (\p\pz. w p * ci p)" unfolding ci_def w_def proof (rule Cauchy_theorem_singularities[OF \open s\ \connected s\ finite _ \valid_path g\ loop path_img homo]) have "open (s - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto then show "ff holomorphic_on s - pz" unfolding ff_def using f_holo h_holo by (auto intro!: holomorphic_intros simp add:pz_def) next show "\p\s. 0 < get_e p \ (\w\cball p (get_e p). w \ s \ (w \ p \ w \ pz))" using get_e using avoid_def by blast qed also have "... = (\p\pz. c * w p * h p * zorder f p)" proof (rule sum.cong[of pz pz,simplified]) fix p assume "p \ pz" show "w p * ci p = c * w p * h p * (zorder f p)" proof (cases "p\s") assume "p \ s" have "ci p = c * h p * (zorder f p)" unfolding ci_def apply (rule contour_integral_unique) using get_e \p\s\ \p\pz\ unfolding cont_def by (metis mult.assoc mult.commute) thus ?thesis by auto next assume "p\s" then have "w p=0" using homo unfolding w_def by auto then show ?thesis by auto qed qed also have "... = c*(\p\pz. w p * h p * zorder f p)" unfolding sum_distrib_left by (simp add:algebra_simps) finally have "contour_integral g ff = c * (\p\pz. w p * h p * of_int (zorder f p))" . then show ?thesis unfolding ff_def c_def w_def by simp qed subsection \Rouche's theorem \ theorem Rouche_theorem: fixes f g::"complex \ complex" and s:: "complex set" defines "fg\(\p. f p + g p)" defines "zeros_fg\{p. fg p = 0}" and "zeros_f\{p. f p = 0}" assumes "open s" and "connected s" and "finite zeros_fg" and "finite zeros_f" and f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" and "valid_path \" and loop:"pathfinish \ = pathstart \" and path_img:"path_image \ \ s " and path_less:"\z\path_image \. cmod(f z) > cmod(g z)" and homo:"\z. (z \ s) \ winding_number \ z = 0" shows "(\p\zeros_fg. winding_number \ p * zorder fg p) = (\p\zeros_f. winding_number \ p * zorder f p)" proof - have path_fg:"path_image \ \ s - zeros_fg" proof - have False when "z\path_image \" and "f z + g z=0" for z proof - have "cmod (f z) > cmod (g z)" using \z\path_image \\ path_less by auto moreover have "f z = - g z" using \f z + g z =0\ by (simp add: eq_neg_iff_add_eq_0) then have "cmod (f z) = cmod (g z)" by auto ultimately show False by auto qed then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto qed have path_f:"path_image \ \ s - zeros_f" proof - have False when "z\path_image \" and "f z =0" for z proof - have "cmod (g z) < cmod (f z) " using \z\path_image \\ path_less by auto then have "cmod (g z) < 0" using \f z=0\ by auto then show False by auto qed then show ?thesis unfolding zeros_f_def using path_img by auto qed define w where "w \ \p. winding_number \ p" define c where "c \ 2 * complex_of_real pi * \" define h where "h \ \p. g p / f p + 1" obtain spikes where "finite spikes" and spikes: "\x\{0..1} - spikes. \ differentiable at x" using \valid_path \\ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have h_contour:"((\x. deriv h x / h x) has_contour_integral 0) \" proof - have outside_img:"0 \ outside (path_image (h o \))" proof - have "h p \ ball 1 1" when "p\path_image \" for p proof - have "cmod (g p/f p) <1" using path_less[rule_format,OF that] apply (cases "cmod (f p) = 0") by (auto simp add: norm_divide) then show ?thesis unfolding h_def by (auto simp add:dist_complex_def) qed then have "path_image (h o \) \ ball 1 1" by (simp add: image_subset_iff path_image_compose) moreover have " (0::complex) \ ball 1 1" by (simp add: dist_norm) ultimately show "?thesis" using convex_in_outside[of "ball 1 1" 0] outside_mono by blast qed have valid_h:"valid_path (h \ \)" proof (rule valid_path_compose_holomorphic[OF \valid_path \\ _ _ path_f]) show "h holomorphic_on s - zeros_f" unfolding h_def using f_holo g_holo by (auto intro!: holomorphic_intros simp add:zeros_f_def) next show "open (s - zeros_f)" using \finite zeros_f\ \open s\ finite_imp_closed by auto qed have "((\z. 1/z) has_contour_integral 0) (h \ \)" proof - have "0 \ path_image (h \ \)" using outside_img by (simp add: outside_def) then have "((\z. 1/z) has_contour_integral c * winding_number (h \ \) 0) (h \ \)" using has_contour_integral_winding_number[of "h o \" 0,simplified] valid_h unfolding c_def by auto moreover have "winding_number (h o \) 0 = 0" proof - have "0 \ outside (path_image (h \ \))" using outside_img . moreover have "path (h o \)" using valid_h by (simp add: valid_path_imp_path) moreover have "pathfinish (h o \) = pathstart (h o \)" by (simp add: loop pathfinish_compose pathstart_compose) ultimately show ?thesis using winding_number_zero_in_outside by auto qed ultimately show ?thesis by auto qed moreover have "vector_derivative (h \ \) (at x) = vector_derivative \ (at x) * deriv h (\ x)" when "x\{0..1} - spikes" for x proof (rule vector_derivative_chain_at_general) show "\ differentiable at x" using that \valid_path \\ spikes by auto next define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" define t where "t \ \ x" have "f t\0" unfolding zeros_f_def t_def by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) moreover have "t\s" using contra_subsetD path_image_def path_fg t_def that by fastforce ultimately have "(h has_field_derivative der t) (at t)" unfolding h_def der_def using g_holo f_holo \open s\ by (auto intro!: holomorphic_derivI derivative_eq_intros) then show "h field_differentiable at (\ x)" unfolding t_def field_differentiable_def by blast qed then have " ((/) 1 has_contour_integral 0) (h \ \) = ((\x. deriv h x / h x) has_contour_integral 0) \" unfolding has_contour_integral apply (intro has_integral_spike_eq[OF negligible_finite, OF \finite spikes\]) by auto ultimately show ?thesis by auto qed then have "contour_integral \ (\x. deriv h x / h x) = 0" using contour_integral_unique by simp moreover have "contour_integral \ (\x. deriv fg x / fg x) = contour_integral \ (\x. deriv f x / f x) + contour_integral \ (\p. deriv h p / h p)" proof - have "(\p. deriv f p / f p) contour_integrable_on \" proof (rule contour_integrable_holomorphic_simple[OF _ _ \valid_path \\ path_f]) show "open (s - zeros_f)" using finite_imp_closed[OF \finite zeros_f\] \open s\ by auto then show "(\p. deriv f p / f p) holomorphic_on s - zeros_f" using f_holo by (auto intro!: holomorphic_intros simp add:zeros_f_def) qed moreover have "(\p. deriv h p / h p) contour_integrable_on \" using h_contour by (simp add: has_contour_integral_integrable) ultimately have "contour_integral \ (\x. deriv f x / f x + deriv h x / h x) = contour_integral \ (\p. deriv f p / f p) + contour_integral \ (\p. deriv h p / h p)" using contour_integral_add[of "(\p. deriv f p / f p)" \ "(\p. deriv h p / h p)" ] by auto moreover have "deriv fg p / fg p = deriv f p / f p + deriv h p / h p" when "p\ path_image \" for p proof - have "fg p\0" and "f p\0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def by auto have "h p\0" proof (rule ccontr) assume "\ h p \ 0" then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) then have "cmod (g p/f p) = 1" by auto moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] apply (cases "cmod (f p) = 0") by (auto simp add: norm_divide) ultimately show False by auto qed have der_fg:"deriv fg p = deriv f p + deriv g p" unfolding fg_def using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _ \open s\] path_img that by auto have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" proof - define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" have "p\s" using path_img that by auto then have "(h has_field_derivative der p) (at p)" unfolding h_def der_def using g_holo f_holo \open s\ \f p\0\ by (auto intro!: derivative_eq_intros holomorphic_derivI) then show ?thesis unfolding der_def using DERIV_imp_deriv by auto qed show ?thesis apply (simp only:der_fg der_h) apply (auto simp add:field_simps \h p\0\ \f p\0\ \fg p\0\) by (auto simp add:field_simps h_def \f p\0\ fg_def) qed then have "contour_integral \ (\p. deriv fg p / fg p) = contour_integral \ (\p. deriv f p / f p + deriv h p / h p)" by (elim contour_integral_eq) ultimately show ?thesis by auto qed moreover have "contour_integral \ (\x. deriv fg x / fg x) = c * (\p\zeros_fg. w p * zorder fg p)" unfolding c_def zeros_fg_def w_def proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo , of _ "{}" "\_. 1",simplified]) show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto show "path_image \ \ s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def . show " finite {p. fg p = 0}" using \finite zeros_fg\ unfolding zeros_fg_def . qed moreover have "contour_integral \ (\x. deriv f x / f x) = c * (\p\zeros_f. w p * zorder f p)" unfolding c_def zeros_f_def w_def proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo , of _ "{}" "\_. 1",simplified]) show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto show "path_image \ \ s - {p. f p = 0}" using path_f unfolding zeros_f_def . show " finite {p. f p = 0}" using \finite zeros_f\ unfolding zeros_f_def . qed ultimately have " c* (\p\zeros_fg. w p * (zorder fg p)) = c* (\p\zeros_f. w p * (zorder f p))" by auto then show ?thesis unfolding c_def using w_def by auto qed end diff --git a/src/HOL/Analysis/Contour_Integration.thy b/src/HOL/Analysis/Contour_Integration.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Analysis/Contour_Integration.thy @@ -0,0 +1,2681 @@ +section \Contour Integration\ + +theory Contour_Integration + imports Henstock_Kurzweil_Integration Path_Connected Complex_Transcendental +begin + +subsection\<^marker>\tag unimportant\ \Piecewise differentiable functions\ + +definition piecewise_differentiable_on + (infixr "piecewise'_differentiable'_on" 50) + where "f piecewise_differentiable_on i \ + continuous_on i f \ + (\S. finite S \ (\x \ i - S. f differentiable (at x within i)))" + +lemma piecewise_differentiable_on_imp_continuous_on: + "f piecewise_differentiable_on S \ continuous_on S f" +by (simp add: piecewise_differentiable_on_def) + +lemma piecewise_differentiable_on_subset: + "f piecewise_differentiable_on S \ T \ S \ f piecewise_differentiable_on T" + using continuous_on_subset + unfolding piecewise_differentiable_on_def + apply safe + apply (blast elim: continuous_on_subset) + by (meson Diff_iff differentiable_within_subset subsetCE) + +lemma differentiable_on_imp_piecewise_differentiable: + fixes a:: "'a::{linorder_topology,real_normed_vector}" + shows "f differentiable_on {a..b} \ f piecewise_differentiable_on {a..b}" + apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on) + apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def) + done + +lemma differentiable_imp_piecewise_differentiable: + "(\x. x \ S \ f differentiable (at x within S)) + \ f piecewise_differentiable_on S" +by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def + intro: differentiable_within_subset) + +lemma piecewise_differentiable_const [iff]: "(\x. z) piecewise_differentiable_on S" + by (simp add: differentiable_imp_piecewise_differentiable) + +lemma piecewise_differentiable_compose: + "\f piecewise_differentiable_on S; g piecewise_differentiable_on (f ` S); + \x. finite (S \ f-`{x})\ + \ (g \ f) piecewise_differentiable_on S" + apply (simp add: piecewise_differentiable_on_def, safe) + apply (blast intro: continuous_on_compose2) + apply (rename_tac A B) + apply (rule_tac x="A \ (\x\B. S \ f-`{x})" in exI) + apply (blast intro!: differentiable_chain_within) + done + +lemma piecewise_differentiable_affine: + fixes m::real + assumes "f piecewise_differentiable_on ((\x. m *\<^sub>R x + c) ` S)" + shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_differentiable_on S" +proof (cases "m = 0") + case True + then show ?thesis + unfolding o_def + by (force intro: differentiable_imp_piecewise_differentiable differentiable_const) +next + case False + show ?thesis + apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable]) + apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+ + done +qed + +lemma piecewise_differentiable_cases: + fixes c::real + assumes "f piecewise_differentiable_on {a..c}" + "g piecewise_differentiable_on {c..b}" + "a \ c" "c \ b" "f c = g c" + shows "(\x. if x \ c then f x else g x) piecewise_differentiable_on {a..b}" +proof - + obtain S T where st: "finite S" "finite T" + and fd: "\x. x \ {a..c} - S \ f differentiable at x within {a..c}" + and gd: "\x. x \ {c..b} - T \ g differentiable at x within {c..b}" + using assms + by (auto simp: piecewise_differentiable_on_def) + have finabc: "finite ({a,b,c} \ (S \ T))" + by (metis \finite S\ \finite T\ finite_Un finite_insert finite.emptyI) + have "continuous_on {a..c} f" "continuous_on {c..b} g" + using assms piecewise_differentiable_on_def by auto + then have "continuous_on {a..b} (\x. if x \ c then f x else g x)" + using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], + OF closed_real_atLeastAtMost [of c b], + of f g "\x. x\c"] assms + by (force simp: ivl_disj_un_two_touch) + moreover + { fix x + assume x: "x \ {a..b} - ({a,b,c} \ (S \ T))" + have "(\x. if x \ c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg") + proof (cases x c rule: le_cases) + case le show ?diff_fg + proof (rule differentiable_transform_within [where d = "dist x c"]) + have "f differentiable at x" + using x le fd [of x] at_within_interior [of x "{a..c}"] by simp + then show "f differentiable at x within {a..b}" + by (simp add: differentiable_at_withinI) + qed (use x le st dist_real_def in auto) + next + case ge show ?diff_fg + proof (rule differentiable_transform_within [where d = "dist x c"]) + have "g differentiable at x" + using x ge gd [of x] at_within_interior [of x "{c..b}"] by simp + then show "g differentiable at x within {a..b}" + by (simp add: differentiable_at_withinI) + qed (use x ge st dist_real_def in auto) + qed + } + then have "\S. finite S \ + (\x\{a..b} - S. (\x. if x \ c then f x else g x) differentiable at x within {a..b})" + by (meson finabc) + ultimately show ?thesis + by (simp add: piecewise_differentiable_on_def) +qed + +lemma piecewise_differentiable_neg: + "f piecewise_differentiable_on S \ (\x. -(f x)) piecewise_differentiable_on S" + by (auto simp: piecewise_differentiable_on_def continuous_on_minus) + +lemma piecewise_differentiable_add: + assumes "f piecewise_differentiable_on i" + "g piecewise_differentiable_on i" + shows "(\x. f x + g x) piecewise_differentiable_on i" +proof - + obtain S T where st: "finite S" "finite T" + "\x\i - S. f differentiable at x within i" + "\x\i - T. g differentiable at x within i" + using assms by (auto simp: piecewise_differentiable_on_def) + then have "finite (S \ T) \ (\x\i - (S \ T). (\x. f x + g x) differentiable at x within i)" + by auto + moreover have "continuous_on i f" "continuous_on i g" + using assms piecewise_differentiable_on_def by auto + ultimately show ?thesis + by (auto simp: piecewise_differentiable_on_def continuous_on_add) +qed + +lemma piecewise_differentiable_diff: + "\f piecewise_differentiable_on S; g piecewise_differentiable_on S\ + \ (\x. f x - g x) piecewise_differentiable_on S" + unfolding diff_conv_add_uminus + by (metis piecewise_differentiable_add piecewise_differentiable_neg) + +lemma continuous_on_joinpaths_D1: + "continuous_on {0..1} (g1 +++ g2) \ continuous_on {0..1} g1" + apply (rule continuous_on_eq [of _ "(g1 +++ g2) \ ((*)(inverse 2))"]) + apply (rule continuous_intros | simp)+ + apply (auto elim!: continuous_on_subset simp: joinpaths_def) + done + +lemma continuous_on_joinpaths_D2: + "\continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\ \ continuous_on {0..1} g2" + apply (rule continuous_on_eq [of _ "(g1 +++ g2) \ (\x. inverse 2*x + 1/2)"]) + apply (rule continuous_intros | simp)+ + apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def) + done + +lemma piecewise_differentiable_D1: + assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" + shows "g1 piecewise_differentiable_on {0..1}" +proof - + obtain S where cont: "continuous_on {0..1} g1" and "finite S" + and S: "\x. x \ {0..1} - S \ g1 +++ g2 differentiable at x within {0..1}" + using assms unfolding piecewise_differentiable_on_def + by (blast dest!: continuous_on_joinpaths_D1) + show ?thesis + unfolding piecewise_differentiable_on_def + proof (intro exI conjI ballI cont) + show "finite (insert 1 (((*)2) ` S))" + by (simp add: \finite S\) + show "g1 differentiable at x within {0..1}" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x + proof (rule_tac d="dist (x/2) (1/2)" in differentiable_transform_within) + have "g1 +++ g2 differentiable at (x / 2) within {0..1/2}" + by (rule differentiable_subset [OF S [of "x/2"]] | use that in force)+ + then show "g1 +++ g2 \ (*) (inverse 2) differentiable at x within {0..1}" + using image_affinity_atLeastAtMost_div [of 2 0 "0::real" 1] + by (auto intro: differentiable_chain_within) + qed (use that in \auto simp: joinpaths_def\) + qed +qed + +lemma piecewise_differentiable_D2: + assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" and eq: "pathfinish g1 = pathstart g2" + shows "g2 piecewise_differentiable_on {0..1}" +proof - + have [simp]: "g1 1 = g2 0" + using eq by (simp add: pathfinish_def pathstart_def) + obtain S where cont: "continuous_on {0..1} g2" and "finite S" + and S: "\x. x \ {0..1} - S \ g1 +++ g2 differentiable at x within {0..1}" + using assms unfolding piecewise_differentiable_on_def + by (blast dest!: continuous_on_joinpaths_D2) + show ?thesis + unfolding piecewise_differentiable_on_def + proof (intro exI conjI ballI cont) + show "finite (insert 0 ((\x. 2*x-1)`S))" + by (simp add: \finite S\) + show "g2 differentiable at x within {0..1}" if "x \ {0..1} - insert 0 ((\x. 2*x-1)`S)" for x + proof (rule_tac d="dist ((x+1)/2) (1/2)" in differentiable_transform_within) + have x2: "(x + 1) / 2 \ S" + using that + apply (clarsimp simp: image_iff) + by (metis add.commute add_diff_cancel_left' mult_2 field_sum_of_halves) + have "g1 +++ g2 \ (\x. (x+1) / 2) differentiable at x within {0..1}" + by (rule differentiable_chain_within differentiable_subset [OF S [of "(x+1)/2"]] | use x2 that in force)+ + then show "g1 +++ g2 \ (\x. (x+1) / 2) differentiable at x within {0..1}" + by (auto intro: differentiable_chain_within) + show "(g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" if "x' \ {0..1}" "dist x' x < dist ((x + 1) / 2) (1/2)" for x' + proof - + have [simp]: "(2*x'+2)/2 = x'+1" + by (simp add: field_split_simps) + show ?thesis + using that by (auto simp: joinpaths_def) + qed + qed (use that in \auto simp: joinpaths_def\) + qed +qed + + +subsection\The concept of continuously differentiable\ + +text \ +John Harrison writes as follows: + +``The usual assumption in complex analysis texts is that a path \\\ should be piecewise +continuously differentiable, which ensures that the path integral exists at least for any continuous +f, since all piecewise continuous functions are integrable. However, our notion of validity is +weaker, just piecewise differentiability\ldots{} [namely] continuity plus differentiability except on a +finite set\ldots{} [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to +the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this +can integrate all derivatives.'' + +"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec. +Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165. + +And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably +difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem +asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ + +definition\<^marker>\tag important\ C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" + (infix "C1'_differentiable'_on" 50) + where + "f C1_differentiable_on S \ + (\D. (\x \ S. (f has_vector_derivative (D x)) (at x)) \ continuous_on S D)" + +lemma C1_differentiable_on_eq: + "f C1_differentiable_on S \ + (\x \ S. f differentiable at x) \ continuous_on S (\x. vector_derivative f (at x))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + unfolding C1_differentiable_on_def + by (metis (no_types, lifting) continuous_on_eq differentiableI_vector vector_derivative_at) +next + assume ?rhs + then show ?lhs + using C1_differentiable_on_def vector_derivative_works by fastforce +qed + +lemma C1_differentiable_on_subset: + "f C1_differentiable_on T \ S \ T \ f C1_differentiable_on S" + unfolding C1_differentiable_on_def continuous_on_eq_continuous_within + by (blast intro: continuous_within_subset) + +lemma C1_differentiable_compose: + assumes fg: "f C1_differentiable_on S" "g C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" + shows "(g \ f) C1_differentiable_on S" +proof - + have "\x. x \ S \ g \ f differentiable at x" + by (meson C1_differentiable_on_eq assms differentiable_chain_at imageI) + moreover have "continuous_on S (\x. vector_derivative (g \ f) (at x))" + proof (rule continuous_on_eq [of _ "\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) + show "continuous_on S (\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)))" + using fg + apply (clarsimp simp add: C1_differentiable_on_eq) + apply (rule Limits.continuous_on_scaleR, assumption) + by (metis (mono_tags, lifting) continuous_at_imp_continuous_on continuous_on_compose continuous_on_cong differentiable_imp_continuous_within o_def) + show "\x. x \ S \ vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)) = vector_derivative (g \ f) (at x)" + by (metis (mono_tags, hide_lams) C1_differentiable_on_eq fg imageI vector_derivative_chain_at) + qed + ultimately show ?thesis + by (simp add: C1_differentiable_on_eq) +qed + +lemma C1_diff_imp_diff: "f C1_differentiable_on S \ f differentiable_on S" + by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on) + +lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\x. x) C1_differentiable_on S" + by (auto simp: C1_differentiable_on_eq) + +lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" + by (auto simp: C1_differentiable_on_eq) + +lemma C1_differentiable_on_add [simp, derivative_intros]: + "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x + g x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_minus [simp, derivative_intros]: + "f C1_differentiable_on S \ (\x. - f x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_diff [simp, derivative_intros]: + "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x - g x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_mult [simp, derivative_intros]: + fixes f g :: "real \ 'a :: real_normed_algebra" + shows "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x * g x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq + by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) + +lemma C1_differentiable_on_scaleR [simp, derivative_intros]: + "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x *\<^sub>R g x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq + by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ + + +definition\<^marker>\tag important\ piecewise_C1_differentiable_on + (infixr "piecewise'_C1'_differentiable'_on" 50) + where "f piecewise_C1_differentiable_on i \ + continuous_on i f \ + (\S. finite S \ (f C1_differentiable_on (i - S)))" + +lemma C1_differentiable_imp_piecewise: + "f C1_differentiable_on S \ f piecewise_C1_differentiable_on S" + by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) + +lemma piecewise_C1_imp_differentiable: + "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" + by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def + C1_differentiable_on_def differentiable_def has_vector_derivative_def + intro: has_derivative_at_withinI) + +lemma piecewise_C1_differentiable_compose: + assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" + shows "(g \ f) piecewise_C1_differentiable_on S" +proof - + have "continuous_on S (\x. g (f x))" + by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def) + moreover have "\T. finite T \ g \ f C1_differentiable_on S - T" + proof - + obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S" + using fg by (auto simp: piecewise_C1_differentiable_on_def) + obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S" + using fg by (auto simp: piecewise_C1_differentiable_on_def) + show ?thesis + proof (intro exI conjI) + show "finite (F \ (\x\G. S \ f-`{x}))" + using fin by (auto simp only: Int_Union \finite F\ \finite G\ finite_UN finite_imageI) + show "g \ f C1_differentiable_on S - (F \ (\x\G. S \ f -` {x}))" + apply (rule C1_differentiable_compose) + apply (blast intro: C1_differentiable_on_subset [OF F]) + apply (blast intro: C1_differentiable_on_subset [OF G]) + by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin) + qed + qed + ultimately show ?thesis + by (simp add: piecewise_C1_differentiable_on_def) +qed + +lemma piecewise_C1_differentiable_on_subset: + "f piecewise_C1_differentiable_on S \ T \ S \ f piecewise_C1_differentiable_on T" + by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) + +lemma C1_differentiable_imp_continuous_on: + "f C1_differentiable_on S \ continuous_on S f" + unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within + using differentiable_at_withinI differentiable_imp_continuous_within by blast + +lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" + unfolding C1_differentiable_on_def + by auto + +lemma piecewise_C1_differentiable_affine: + fixes m::real + assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` S)" + shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on S" +proof (cases "m = 0") + case True + then show ?thesis + unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def) +next + case False + have *: "\x. finite (S \ {y. m * y + c = x})" + using False not_finite_existsD by fastforce + show ?thesis + apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) + apply (rule * assms derivative_intros | simp add: False vimage_def)+ + done +qed + +lemma piecewise_C1_differentiable_cases: + fixes c::real + assumes "f piecewise_C1_differentiable_on {a..c}" + "g piecewise_C1_differentiable_on {c..b}" + "a \ c" "c \ b" "f c = g c" + shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" +proof - + obtain S T where st: "f C1_differentiable_on ({a..c} - S)" + "g C1_differentiable_on ({c..b} - T)" + "finite S" "finite T" + using assms + by (force simp: piecewise_C1_differentiable_on_def) + then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" + using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], + OF closed_real_atLeastAtMost [of c b], + of f g "\x. x\c"] assms + by (force simp: ivl_disj_un_two_touch) + { fix x + assume x: "x \ {a..b} - insert c (S \ T)" + have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") + proof (cases x c rule: le_cases) + case le show ?diff_fg + apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) + using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) + next + case ge show ?diff_fg + apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) + using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) + qed + } + then have "(\x \ {a..b} - insert c (S \ T). (\x. if x \ c then f x else g x) differentiable at x)" + by auto + moreover + { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" + and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" + have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + proof - + have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)" + if "a < x" "x < c" "x \ S" for x + proof - + have f: "f differentiable at x" + by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that) + show ?thesis + using that + apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) + apply (auto simp: dist_norm vector_derivative_works [symmetric] f) + done + qed + then show ?thesis + by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at) + qed + moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + proof - + have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)" + if "c < x" "x < b" "x \ T" for x + proof - + have g: "g differentiable at x" + by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that) + show ?thesis + using that + apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) + apply (auto simp: dist_norm vector_derivative_works [symmetric] g) + done + qed + then show ?thesis + by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at) + qed + ultimately have "continuous_on ({a<.. T)) + (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + by (rule continuous_on_subset [OF continuous_on_open_Un], auto) + } note * = this + have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + using st + by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) + ultimately have "\S. finite S \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - S)" + apply (rule_tac x="{a,b,c} \ S \ T" in exI) + using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) + with cab show ?thesis + by (simp add: piecewise_C1_differentiable_on_def) +qed + +lemma piecewise_C1_differentiable_neg: + "f piecewise_C1_differentiable_on S \ (\x. -(f x)) piecewise_C1_differentiable_on S" + unfolding piecewise_C1_differentiable_on_def + by (auto intro!: continuous_on_minus C1_differentiable_on_minus) + +lemma piecewise_C1_differentiable_add: + assumes "f piecewise_C1_differentiable_on i" + "g piecewise_C1_differentiable_on i" + shows "(\x. f x + g x) piecewise_C1_differentiable_on i" +proof - + obtain S t where st: "finite S" "finite t" + "f C1_differentiable_on (i-S)" + "g C1_differentiable_on (i-t)" + using assms by (auto simp: piecewise_C1_differentiable_on_def) + then have "finite (S \ t) \ (\x. f x + g x) C1_differentiable_on i - (S \ t)" + by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) + moreover have "continuous_on i f" "continuous_on i g" + using assms piecewise_C1_differentiable_on_def by auto + ultimately show ?thesis + by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) +qed + +lemma piecewise_C1_differentiable_diff: + "\f piecewise_C1_differentiable_on S; g piecewise_C1_differentiable_on S\ + \ (\x. f x - g x) piecewise_C1_differentiable_on S" + unfolding diff_conv_add_uminus + by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) + +lemma piecewise_C1_differentiable_D1: + fixes g1 :: "real \ 'a::real_normed_field" + assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" + shows "g1 piecewise_C1_differentiable_on {0..1}" +proof - + obtain S where "finite S" + and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" + and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" + using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have g1D: "g1 differentiable at x" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x + proof (rule differentiable_transform_within) + show "g1 +++ g2 \ (*) (inverse 2) differentiable at x" + using that g12D + apply (simp only: joinpaths_def) + by (rule differentiable_chain_at derivative_intros | force)+ + show "\x'. \dist x' x < dist (x/2) (1/2)\ + \ (g1 +++ g2 \ (*) (inverse 2)) x' = g1 x'" + using that by (auto simp: dist_real_def joinpaths_def) + qed (use that in \auto simp: dist_real_def\) + have [simp]: "vector_derivative (g1 \ (*) 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" + if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x + apply (subst vector_derivative_chain_at) + using that + apply (rule derivative_eq_intros g1D | simp)+ + done + have "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" + using co12 by (rule continuous_on_subset) force + then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 \ (*)2) (at x))" + proof (rule continuous_on_eq [OF _ vector_derivative_at]) + show "(g1 +++ g2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" + if "x \ {0..1/2} - insert (1/2) S" for x + proof (rule has_vector_derivative_transform_within) + show "(g1 \ (*) 2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" + using that + by (force intro: g1D differentiable_chain_at simp: vector_derivative_works [symmetric]) + show "\x'. \dist x' x < dist x (1/2)\ \ (g1 \ (*) 2) x' = (g1 +++ g2) x'" + using that by (auto simp: dist_norm joinpaths_def) + qed (use that in \auto simp: dist_norm\) + qed + have "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) + ((\x. 1/2 * vector_derivative (g1 \ (*)2) (at x)) \ (*)(1/2))" + apply (rule continuous_intros)+ + using coDhalf + apply (simp add: scaleR_conv_of_real image_set_diff image_image) + done + then have con_g1: "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) (\x. vector_derivative g1 (at x))" + by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) + have "continuous_on {0..1} g1" + using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast + with \finite S\ show ?thesis + apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + apply (rule_tac x="insert 1 (((*)2)`S)" in exI) + apply (simp add: g1D con_g1) + done +qed + +lemma piecewise_C1_differentiable_D2: + fixes g2 :: "real \ 'a::real_normed_field" + assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" + shows "g2 piecewise_C1_differentiable_on {0..1}" +proof - + obtain S where "finite S" + and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" + and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" + using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have g2D: "g2 differentiable at x" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x + proof (rule differentiable_transform_within) + show "g1 +++ g2 \ (\x. (x + 1) / 2) differentiable at x" + using g12D that + apply (simp only: joinpaths_def) + apply (drule_tac x= "(x+1) / 2" in bspec, force simp: field_split_simps) + apply (rule differentiable_chain_at derivative_intros | force)+ + done + show "\x'. dist x' x < dist ((x + 1) / 2) (1/2) \ (g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" + using that by (auto simp: dist_real_def joinpaths_def field_simps) + qed (use that in \auto simp: dist_norm\) + have [simp]: "vector_derivative (g2 \ (\x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" + if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x + using that by (auto simp: vector_derivative_chain_at field_split_simps g2D) + have "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" + using co12 by (rule continuous_on_subset) force + then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g2 \ (\x. 2*x-1)) (at x))" + proof (rule continuous_on_eq [OF _ vector_derivative_at]) + show "(g1 +++ g2 has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) + (at x)" + if "x \ {1 / 2..1} - insert (1 / 2) S" for x + proof (rule_tac f="g2 \ (\x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) + show "(g2 \ (\x. 2 * x - 1) has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) + (at x)" + using that by (force intro: g2D differentiable_chain_at simp: vector_derivative_works [symmetric]) + show "\x'. \dist x' x < dist (3 / 4) ((x + 1) / 2)\ \ (g2 \ (\x. 2 * x - 1)) x' = (g1 +++ g2) x'" + using that by (auto simp: dist_norm joinpaths_def add_divide_distrib) + qed (use that in \auto simp: dist_norm\) + qed + have [simp]: "((\x. (x+1) / 2) ` ({0..1} - insert 0 ((\x. 2 * x - 1) ` S))) = ({1/2..1} - insert (1/2) S)" + apply (simp add: image_set_diff inj_on_def image_image) + apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) + done + have "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) + ((\x. 1/2 * vector_derivative (g2 \ (\x. 2*x-1)) (at x)) \ (\x. (x+1)/2))" + by (rule continuous_intros | simp add: coDhalf)+ + then have con_g2: "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) (\x. vector_derivative g2 (at x))" + by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) + have "continuous_on {0..1} g2" + using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast + with \finite S\ show ?thesis + apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + apply (rule_tac x="insert 0 ((\x. 2 * x - 1) ` S)" in exI) + apply (simp add: g2D con_g2) + done +qed + +subsection \Valid paths, and their start and finish\ + +definition\<^marker>\tag important\ valid_path :: "(real \ 'a :: real_normed_vector) \ bool" + where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" + +definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" + where "closed_path g \ g 0 = g 1" + +text\In particular, all results for paths apply\ + +lemma valid_path_imp_path: "valid_path g \ path g" + by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) + +lemma connected_valid_path_image: "valid_path g \ connected(path_image g)" + by (metis connected_path_image valid_path_imp_path) + +lemma compact_valid_path_image: "valid_path g \ compact(path_image g)" + by (metis compact_path_image valid_path_imp_path) + +lemma bounded_valid_path_image: "valid_path g \ bounded(path_image g)" + by (metis bounded_path_image valid_path_imp_path) + +lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" + by (metis closed_path_image valid_path_imp_path) + +lemma valid_path_compose: + assumes "valid_path g" + and der: "\x. x \ path_image g \ f field_differentiable (at x)" + and con: "continuous_on (path_image g) (deriv f)" + shows "valid_path (f \ g)" +proof - + obtain S where "finite S" and g_diff: "g C1_differentiable_on {0..1} - S" + using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto + have "f \ g differentiable at t" when "t\{0..1} - S" for t + proof (rule differentiable_chain_at) + show "g differentiable at t" using \valid_path g\ + by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - S\ that) + next + have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis + then show "f differentiable at (g t)" + using der[THEN field_differentiable_imp_differentiable] by auto + qed + moreover have "continuous_on ({0..1} - S) (\x. vector_derivative (f \ g) (at x))" + proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], + rule continuous_intros) + show "continuous_on ({0..1} - S) (\x. vector_derivative g (at x))" + using g_diff C1_differentiable_on_eq by auto + next + have "continuous_on {0..1} (\x. deriv f (g x))" + using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] + \valid_path g\ piecewise_C1_differentiable_on_def valid_path_def + by blast + then show "continuous_on ({0..1} - S) (\x. deriv f (g x))" + using continuous_on_subset by blast + next + show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" + when "t \ {0..1} - S" for t + proof (rule vector_derivative_chain_at_general[symmetric]) + show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) + next + have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis + then show "f field_differentiable at (g t)" using der by auto + qed + qed + ultimately have "f \ g C1_differentiable_on {0..1} - S" + using C1_differentiable_on_eq by blast + moreover have "path (f \ g)" + apply (rule path_continuous_image[OF valid_path_imp_path[OF \valid_path g\]]) + using der + by (simp add: continuous_at_imp_continuous_on field_differentiable_imp_continuous_at) + ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def + using \finite S\ by auto +qed + +lemma valid_path_uminus_comp[simp]: + fixes g::"real \ 'a ::real_normed_field" + shows "valid_path (uminus \ g) \ valid_path g" +proof + show "valid_path g \ valid_path (uminus \ g)" for g::"real \ 'a" + by (auto intro!: valid_path_compose derivative_intros simp add: deriv_linear[of "-1",simplified]) + then show "valid_path g" when "valid_path (uminus \ g)" + by (metis fun.map_comp group_add_class.minus_comp_minus id_comp that) +qed + +lemma valid_path_offset[simp]: + shows "valid_path (\t. g t - z) \ valid_path g" +proof + show *: "valid_path (g::real\'a) \ valid_path (\t. g t - z)" for g z + unfolding valid_path_def + by (fastforce intro:derivative_intros C1_differentiable_imp_piecewise piecewise_C1_differentiable_diff) + show "valid_path (\t. g t - z) \ valid_path g" + using *[of "\t. g t - z" "-z",simplified] . +qed + + +subsection\Contour Integrals along a path\ + +text\This definition is for complex numbers only, and does not generalise to line integrals in a vector field\ + +text\piecewise differentiable function on [0,1]\ + +definition\<^marker>\tag important\ has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" + (infixr "has'_contour'_integral" 50) + where "(f has_contour_integral i) g \ + ((\x. f(g x) * vector_derivative g (at x within {0..1})) + has_integral i) {0..1}" + +definition\<^marker>\tag important\ contour_integrable_on + (infixr "contour'_integrable'_on" 50) + where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" + +definition\<^marker>\tag important\ contour_integral + where "contour_integral g f \ SOME i. (f has_contour_integral i) g \ \ f contour_integrable_on g \ i=0" + +lemma not_integrable_contour_integral: "\ f contour_integrable_on g \ contour_integral g f = 0" + unfolding contour_integrable_on_def contour_integral_def by blast + +lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" + apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) + using has_integral_unique by blast + +lemma has_contour_integral_eqpath: + "\(f has_contour_integral y) p; f contour_integrable_on \; + contour_integral p f = contour_integral \ f\ + \ (f has_contour_integral y) \" +using contour_integrable_on_def contour_integral_unique by auto + +lemma has_contour_integral_integral: + "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" + by (metis contour_integral_unique contour_integrable_on_def) + +lemma has_contour_integral_unique: + "(f has_contour_integral i) g \ (f has_contour_integral j) g \ i = j" + using has_integral_unique + by (auto simp: has_contour_integral_def) + +lemma has_contour_integral_integrable: "(f has_contour_integral i) g \ f contour_integrable_on g" + using contour_integrable_on_def by blast + +text\Show that we can forget about the localized derivative.\ + +lemma has_integral_localized_vector_derivative: + "((\x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \ + ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}" +proof - + have *: "{a..b} - {a,b} = interior {a..b}" + by (simp add: atLeastAtMost_diff_ends) + show ?thesis + apply (rule has_integral_spike_eq [of "{a,b}"]) + apply (auto simp: at_within_interior [of _ "{a..b}"]) + done +qed + +lemma integrable_on_localized_vector_derivative: + "(\x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \ + (\x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}" + by (simp add: integrable_on_def has_integral_localized_vector_derivative) + +lemma has_contour_integral: + "(f has_contour_integral i) g \ + ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" + by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) + +lemma contour_integrable_on: + "f contour_integrable_on g \ + (\t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" + by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) + +subsection\<^marker>\tag unimportant\ \Reversing a path\ + +lemma valid_path_imp_reverse: + assumes "valid_path g" + shows "valid_path(reversepath g)" +proof - + obtain S where "finite S" and S: "g C1_differentiable_on ({0..1} - S)" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + then have "finite ((-) 1 ` S)" + by auto + moreover have "(reversepath g C1_differentiable_on ({0..1} - (-) 1 ` S))" + unfolding reversepath_def + apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) + using S + by (force simp: finite_vimageI inj_on_def C1_differentiable_on_eq elim!: continuous_on_subset)+ + ultimately show ?thesis using assms + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) +qed + +lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" + using valid_path_imp_reverse by force + +lemma has_contour_integral_reversepath: + assumes "valid_path g" and f: "(f has_contour_integral i) g" + shows "(f has_contour_integral (-i)) (reversepath g)" +proof - + { fix S x + assume xs: "g C1_differentiable_on ({0..1} - S)" "x \ (-) 1 ` S" "0 \ x" "x \ 1" + have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = + - vector_derivative g (at (1 - x) within {0..1})" + proof - + obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" + using xs + by (force simp: has_vector_derivative_def C1_differentiable_on_def) + have "(g \ (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" + by (intro vector_diff_chain_within has_vector_derivative_at_within [OF f'] derivative_eq_intros | simp)+ + then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" + by (simp add: o_def) + show ?thesis + using xs + by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) + qed + } note * = this + obtain S where S: "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + have "((\x. - (f (g (1 - x)) * vector_derivative g (at (1 - x) within {0..1}))) has_integral -i) + {0..1}" + using has_integral_affinity01 [where m= "-1" and c=1, OF f [unfolded has_contour_integral_def]] + by (simp add: has_integral_neg) + then show ?thesis + using S + apply (clarsimp simp: reversepath_def has_contour_integral_def) + apply (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) + apply (auto simp: *) + done +qed + +lemma contour_integrable_reversepath: + "valid_path g \ f contour_integrable_on g \ f contour_integrable_on (reversepath g)" + using has_contour_integral_reversepath contour_integrable_on_def by blast + +lemma contour_integrable_reversepath_eq: + "valid_path g \ (f contour_integrable_on (reversepath g) \ f contour_integrable_on g)" + using contour_integrable_reversepath valid_path_reversepath by fastforce + +lemma contour_integral_reversepath: + assumes "valid_path g" + shows "contour_integral (reversepath g) f = - (contour_integral g f)" +proof (cases "f contour_integrable_on g") + case True then show ?thesis + by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) +next + case False then have "\ f contour_integrable_on (reversepath g)" + by (simp add: assms contour_integrable_reversepath_eq) + with False show ?thesis by (simp add: not_integrable_contour_integral) +qed + + +subsection\<^marker>\tag unimportant\ \Joining two paths together\ + +lemma valid_path_join: + assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" + shows "valid_path(g1 +++ g2)" +proof - + have "g1 1 = g2 0" + using assms by (auto simp: pathfinish_def pathstart_def) + moreover have "(g1 \ (\x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" + apply (rule piecewise_C1_differentiable_compose) + using assms + apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) + apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) + done + moreover have "(g2 \ (\x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" + apply (rule piecewise_C1_differentiable_compose) + using assms unfolding valid_path_def piecewise_C1_differentiable_on_def + by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI + simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) + ultimately show ?thesis + apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) + apply (rule piecewise_C1_differentiable_cases) + apply (auto simp: o_def) + done +qed + +lemma valid_path_join_D1: + fixes g1 :: "real \ 'a::real_normed_field" + shows "valid_path (g1 +++ g2) \ valid_path g1" + unfolding valid_path_def + by (rule piecewise_C1_differentiable_D1) + +lemma valid_path_join_D2: + fixes g2 :: "real \ 'a::real_normed_field" + shows "\valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\ \ valid_path g2" + unfolding valid_path_def + by (rule piecewise_C1_differentiable_D2) + +lemma valid_path_join_eq [simp]: + fixes g2 :: "real \ 'a::real_normed_field" + shows "pathfinish g1 = pathstart g2 \ (valid_path(g1 +++ g2) \ valid_path g1 \ valid_path g2)" + using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast + +lemma has_contour_integral_join: + assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" + "valid_path g1" "valid_path g2" + shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" +proof - + obtain s1 s2 + where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" + and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" + using assms + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have 1: "((\x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" + and 2: "((\x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" + using assms + by (auto simp: has_contour_integral) + have i1: "((\x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" + and i2: "((\x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" + using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] + has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] + by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) + have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ + vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = + 2 *\<^sub>R vector_derivative g1 (at (z*2))" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\z - 1/2\"]]) + apply (simp_all add: dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) + apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + using s1 + apply (auto simp: algebra_simps vector_derivative_works) + done + have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ + vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = + 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2 (2*x - 1))" and d = "\z - 1/2\"]]) + apply (simp_all add: dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) + apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + using s2 + apply (auto simp: algebra_simps vector_derivative_works) + done + have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" + apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) ((*)2 -` s1)"]) + using s1 + apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) + apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) + done + moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" + apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) + using s2 + apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) + apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) + done + ultimately + show ?thesis + apply (simp add: has_contour_integral) + apply (rule has_integral_combine [where c = "1/2"], auto) + done +qed + +lemma contour_integrable_joinI: + assumes "f contour_integrable_on g1" "f contour_integrable_on g2" + "valid_path g1" "valid_path g2" + shows "f contour_integrable_on (g1 +++ g2)" + using assms + by (meson has_contour_integral_join contour_integrable_on_def) + +lemma contour_integrable_joinD1: + assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" + shows "f contour_integrable_on g1" +proof - + obtain s1 + where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have "(\x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" + using assms + apply (auto simp: contour_integrable_on) + apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) + apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) + done + then have *: "(\x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" + by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) + have g1: "\0 < z; z < 1; z \ s1\ \ + vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = + 2 *\<^sub>R vector_derivative g1 (at z)" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\(z-1)/2\"]]) + apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. x*2" 2 _ g1, simplified o_def]) + using s1 + apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left) + done + show ?thesis + using s1 + apply (auto simp: contour_integrable_on) + apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) + apply (auto simp: joinpaths_def scaleR_conv_of_real g1) + done +qed + +lemma contour_integrable_joinD2: + assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" + shows "f contour_integrable_on g2" +proof - + obtain s2 + where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have "(\x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" + using assms + apply (auto simp: contour_integrable_on) + apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) + apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) + apply (simp add: image_affinity_atLeastAtMost_diff) + done + then have *: "(\x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) + integrable_on {0..1}" + by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) + have g2: "\0 < z; z < 1; z \ s2\ \ + vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = + 2 *\<^sub>R vector_derivative g2 (at z)" for z + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2(2*x-1))" and d = "\z/2\"]]) + apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. x*2-1" 2 _ g2, simplified o_def]) + using s2 + apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left + vector_derivative_works add_divide_distrib) + done + show ?thesis + using s2 + apply (auto simp: contour_integrable_on) + apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) + apply (auto simp: joinpaths_def scaleR_conv_of_real g2) + done +qed + +lemma contour_integrable_join [simp]: + shows + "\valid_path g1; valid_path g2\ + \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" +using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast + +lemma contour_integral_join [simp]: + shows + "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ + \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" + by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) + + +subsection\<^marker>\tag unimportant\ \Shifting the starting point of a (closed) path\ + +lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" + by (auto simp: shiftpath_def) + +lemma valid_path_shiftpath [intro]: + assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "valid_path(shiftpath a g)" + using assms + apply (auto simp: valid_path_def shiftpath_alt_def) + apply (rule piecewise_C1_differentiable_cases) + apply (auto simp: algebra_simps) + apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) + apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) + apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) + apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) + done + +lemma has_contour_integral_shiftpath: + assumes f: "(f has_contour_integral i) g" "valid_path g" + and a: "a \ {0..1}" + shows "(f has_contour_integral i) (shiftpath a g)" +proof - + obtain s + where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have *: "((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" + using assms by (auto simp: has_contour_integral) + then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + + integral {0..a} (\x. f (g x) * vector_derivative g (at x))" + apply (rule has_integral_unique) + apply (subst add.commute) + apply (subst integral_combine) + using assms * integral_unique by auto + { fix x + have "0 \ x \ x + a < 1 \ x \ (\x. x - a) ` s \ + vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" + unfolding shiftpath_def + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x))" and d = "dist(1-a) x"]]) + apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. x+a" 1 _ g, simplified o_def scaleR_one]) + apply (intro derivative_eq_intros | simp)+ + using g + apply (drule_tac x="x+a" in bspec) + using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) + done + } note vd1 = this + { fix x + have "1 < x + a \ x \ 1 \ x \ (\x. x - a + 1) ` s \ + vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" + unfolding shiftpath_def + apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x-1))" and d = "dist (1-a) x"]]) + apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) + apply (rule vector_diff_chain_at [of "\x. x+a-1" 1 _ g, simplified o_def scaleR_one]) + apply (intro derivative_eq_intros | simp)+ + using g + apply (drule_tac x="x+a-1" in bspec) + using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) + done + } note vd2 = this + have va1: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" + using * a by (fastforce intro: integrable_subinterval_real) + have v0a: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" + apply (rule integrable_subinterval_real) + using * a by auto + have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) + has_integral integral {a..1} (\x. f (g x) * vector_derivative g (at x))) {0..1 - a}" + apply (rule has_integral_spike_finite + [where S = "{1-a} \ (\x. x-a) ` s" and f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) + using s apply blast + using a apply (auto simp: algebra_simps vd1) + apply (force simp: shiftpath_def add.commute) + using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]] + apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute) + done + moreover + have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) + has_integral integral {0..a} (\x. f (g x) * vector_derivative g (at x))) {1 - a..1}" + apply (rule has_integral_spike_finite + [where S = "{1-a} \ (\x. x-a+1) ` s" and f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) + using s apply blast + using a apply (auto simp: algebra_simps vd2) + apply (force simp: shiftpath_def add.commute) + using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] + apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified]) + apply (simp add: algebra_simps) + done + ultimately show ?thesis + using a + by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) +qed + +lemma has_contour_integral_shiftpath_D: + assumes "(f has_contour_integral i) (shiftpath a g)" + "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "(f has_contour_integral i) g" +proof - + obtain s + where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + { fix x + assume x: "0 < x" "x < 1" "x \ s" + then have gx: "g differentiable at x" + using g by auto + have "vector_derivative g (at x within {0..1}) = + vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" + apply (rule vector_derivative_at_within_ivl + [OF has_vector_derivative_transform_within_open + [where f = "(shiftpath (1 - a) (shiftpath a g))" and S = "{0<..<1}-s"]]) + using s g assms x + apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath + at_within_interior [of _ "{0..1}"] vector_derivative_works [symmetric]) + apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"]) + apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) + done + } note vd = this + have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" + using assms by (auto intro!: has_contour_integral_shiftpath) + show ?thesis + apply (simp add: has_contour_integral_def) + apply (rule has_integral_spike_finite [of "{0,1} \ s", OF _ _ fi [unfolded has_contour_integral_def]]) + using s assms vd + apply (auto simp: Path_Connected.shiftpath_shiftpath) + done +qed + +lemma has_contour_integral_shiftpath_eq: + assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "(f has_contour_integral i) (shiftpath a g) \ (f has_contour_integral i) g" + using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast + +lemma contour_integrable_on_shiftpath_eq: + assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" +using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto + +lemma contour_integral_shiftpath: + assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" + shows "contour_integral (shiftpath a g) f = contour_integral g f" + using assms + by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) + + +subsection\<^marker>\tag unimportant\ \More about straight-line paths\ + +lemma has_vector_derivative_linepath_within: + "(linepath a b has_vector_derivative (b - a)) (at x within s)" +apply (simp add: linepath_def has_vector_derivative_def algebra_simps) +apply (rule derivative_eq_intros | simp)+ +done + +lemma vector_derivative_linepath_within: + "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" + apply (rule vector_derivative_within_cbox [of 0 "1::real", simplified]) + apply (auto simp: has_vector_derivative_linepath_within) + done + +lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" + by (simp add: has_vector_derivative_linepath_within vector_derivative_at) + +lemma valid_path_linepath [iff]: "valid_path (linepath a b)" + apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) + apply (rule_tac x="{}" in exI) + apply (simp add: differentiable_on_def differentiable_def) + using has_vector_derivative_def has_vector_derivative_linepath_within + apply (fastforce simp add: continuous_on_eq_continuous_within) + done + +lemma has_contour_integral_linepath: + shows "(f has_contour_integral i) (linepath a b) \ + ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" + by (simp add: has_contour_integral) + +lemma linepath_in_path: + shows "x \ {0..1} \ linepath a b x \ closed_segment a b" + by (auto simp: segment linepath_def) + +lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" + by (auto simp: segment linepath_def) + +lemma linepath_in_convex_hull: + fixes x::real + assumes a: "a \ convex hull s" + and b: "b \ convex hull s" + and x: "0\x" "x\1" + shows "linepath a b x \ convex hull s" + apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) + using x + apply (auto simp: linepath_image_01 [symmetric]) + done + +lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" + by (simp add: linepath_def) + +lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" + by (simp add: linepath_def) + +lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" + by (simp add: has_contour_integral_linepath) + +lemma has_contour_integral_trivial_iff [simp]: "(f has_contour_integral i) (linepath a a) \ i=0" + using has_contour_integral_unique by blast + +lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" + using has_contour_integral_trivial contour_integral_unique by blast + +lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" + by (auto simp: linepath_def) + +lemma bounded_linear_linepath: + assumes "bounded_linear f" + shows "f (linepath a b x) = linepath (f a) (f b) x" +proof - + interpret f: bounded_linear f by fact + show ?thesis by (simp add: linepath_def f.add f.scale) +qed + +lemma bounded_linear_linepath': + assumes "bounded_linear f" + shows "f \ linepath a b = linepath (f a) (f b)" + using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) + +lemma cnj_linepath: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" + by (simp add: linepath_def) + +lemma cnj_linepath': "cnj \ linepath a b = linepath (cnj a) (cnj b)" + by (simp add: linepath_def fun_eq_iff) + +subsection\Relation to subpath construction\ + +lemma valid_path_subpath: + fixes g :: "real \ 'a :: real_normed_vector" + assumes "valid_path g" "u \ {0..1}" "v \ {0..1}" + shows "valid_path(subpath u v g)" +proof (cases "v=u") + case True + then show ?thesis + unfolding valid_path_def subpath_def + by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) +next + case False + have "(g \ (\x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" + apply (rule piecewise_C1_differentiable_compose) + apply (simp add: C1_differentiable_imp_piecewise) + apply (simp add: image_affinity_atLeastAtMost) + using assms False + apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) + apply (subst Int_commute) + apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) + done + then show ?thesis + by (auto simp: o_def valid_path_def subpath_def) +qed + +lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" + by (simp add: has_contour_integral subpath_def) + +lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" + using has_contour_integral_subpath_refl contour_integrable_on_def by blast + +lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" + by (simp add: contour_integral_unique) + +lemma has_contour_integral_subpath: + assumes f: "f contour_integrable_on g" and g: "valid_path g" + and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) + (subpath u v g)" +proof (cases "v=u") + case True + then show ?thesis + using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) +next + case False + obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" + using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast + have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) + has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) + {0..1}" + using f uv + apply (simp add: contour_integrable_on subpath_def has_contour_integral) + apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) + apply (simp_all add: has_integral_integral) + apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) + apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) + apply (simp add: divide_simps False) + done + { fix x + have "x \ {0..1} \ + x \ (\t. (v-u) *\<^sub>R t + u) -` s \ + vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" + apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) + apply (intro derivative_eq_intros | simp)+ + apply (cut_tac s [of "(v - u) * x + u"]) + using uv mult_left_le [of x "v-u"] + apply (auto simp: vector_derivative_works) + done + } note vd = this + show ?thesis + apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) + using fs assms + apply (simp add: False subpath_def has_contour_integral) + apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) + apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) + done +qed + +lemma contour_integrable_subpath: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" + shows "f contour_integrable_on (subpath u v g)" + apply (cases u v rule: linorder_class.le_cases) + apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) + apply (subst reversepath_subpath [symmetric]) + apply (rule contour_integrable_reversepath) + using assms apply (blast intro: valid_path_subpath) + apply (simp add: contour_integrable_on_def) + using assms apply (blast intro: has_contour_integral_subpath) + done + +lemma has_integral_contour_integral_subpath: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "(((\x. f(g x) * vector_derivative g (at x))) + has_integral contour_integral (subpath u v g) f) {u..v}" + using assms + apply (auto simp: has_integral_integrable_integral) + apply (rule integrable_on_subcbox [where a=u and b=v and S = "{0..1}", simplified]) + apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) + done + +lemma contour_integral_subcontour_integral: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "contour_integral (subpath u v g) f = + integral {u..v} (\x. f(g x) * vector_derivative g (at x))" + using assms has_contour_integral_subpath contour_integral_unique by blast + +lemma contour_integral_subpath_combine_less: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" + "u {0..1}" "v \ {0..1}" "w \ {0..1}" + shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = + contour_integral (subpath u w g) f" +proof (cases "u\v \ v\w \ u\w") + case True + have *: "subpath v u g = reversepath(subpath u v g) \ + subpath w u g = reversepath(subpath u w g) \ + subpath w v g = reversepath(subpath v w g)" + by (auto simp: reversepath_subpath) + have "u < v \ v < w \ + u < w \ w < v \ + v < u \ u < w \ + v < w \ w < u \ + w < u \ u < v \ + w < v \ v < u" + using True assms by linarith + with assms show ?thesis + using contour_integral_subpath_combine_less [of f g u v w] + contour_integral_subpath_combine_less [of f g u w v] + contour_integral_subpath_combine_less [of f g v u w] + contour_integral_subpath_combine_less [of f g v w u] + contour_integral_subpath_combine_less [of f g w u v] + contour_integral_subpath_combine_less [of f g w v u] + apply simp + apply (elim disjE) + apply (auto simp: * contour_integral_reversepath contour_integrable_subpath + valid_path_subpath algebra_simps) + done +next + case False + then show ?thesis + apply (auto) + using assms + by (metis eq_neg_iff_add_eq_0 contour_integral_reversepath reversepath_subpath valid_path_subpath) +qed + +lemma contour_integral_integral: + "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" + by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) + +lemma contour_integral_cong: + assumes "g = g'" "\x. x \ path_image g \ f x = f' x" + shows "contour_integral g f = contour_integral g' f'" + unfolding contour_integral_integral using assms + by (intro integral_cong) (auto simp: path_image_def) + + +text \Contour integral along a segment on the real axis\ + +lemma has_contour_integral_linepath_Reals_iff: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "(f has_contour_integral I) (linepath a b) \ + ((\x. f (of_real x)) has_integral I) {Re a..Re b}" +proof - + from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" + by (simp_all add: complex_eq_iff) + from assms have "a \ b" by auto + have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ + ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" + by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) + (insert assms, simp_all add: field_simps scaleR_conv_of_real) + also have "(\x. f (a + b * of_real x - a * of_real x)) = + (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" + using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) + also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ + ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms + by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) + also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def + by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) + finally show ?thesis by simp +qed + +lemma contour_integrable_linepath_Reals_iff: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "(f contour_integrable_on linepath a b) \ + (\x. f (of_real x)) integrable_on {Re a..Re b}" + using has_contour_integral_linepath_Reals_iff[OF assms, of f] + by (auto simp: contour_integrable_on_def integrable_on_def) + +lemma contour_integral_linepath_Reals_eq: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" +proof (cases "f contour_integrable_on linepath a b") + case True + thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] + using has_contour_integral_integral has_contour_integral_unique by blast +next + case False + thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] + by (simp add: not_integrable_contour_integral not_integrable_integral) +qed + + + +text\Cauchy's theorem where there's a primitive\ + +lemma contour_integral_primitive_lemma: + fixes f :: "complex \ complex" and g :: "real \ complex" + assumes "a \ b" + and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" + shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) + has_integral (f(g b) - f(g a))) {a..b}" +proof - + obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" + using assms by (auto simp: piecewise_differentiable_on_def) + have cfg: "continuous_on {a..b} (\x. f (g x))" + apply (rule continuous_on_compose [OF cg, unfolded o_def]) + using assms + apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) + done + { fix x::real + assume a: "a < x" and b: "x < b" and xk: "x \ k" + then have "g differentiable at x within {a..b}" + using k by (simp add: differentiable_at_withinI) + then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" + by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) + then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" + by (simp add: has_vector_derivative_def scaleR_conv_of_real) + have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" + using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) + then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" + by (simp add: has_field_derivative_def) + have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" + using diff_chain_within [OF gdiff fdiff] + by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) + } note * = this + show ?thesis + apply (rule fundamental_theorem_of_calculus_interior_strong) + using k assms cfg * + apply (auto simp: at_within_Icc_at) + done +qed + +lemma contour_integral_primitive: + assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "valid_path g" "path_image g \ s" + shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" + using assms + apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) + apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) + done + +corollary Cauchy_theorem_primitive: + assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "valid_path g" "path_image g \ s" "pathfinish g = pathstart g" + shows "(f' has_contour_integral 0) g" + using assms + by (metis diff_self contour_integral_primitive) + +text\Existence of path integral for continuous function\ +lemma contour_integrable_continuous_linepath: + assumes "continuous_on (closed_segment a b) f" + shows "f contour_integrable_on (linepath a b)" +proof - + have "continuous_on {0..1} ((\x. f x * (b - a)) \ linepath a b)" + apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) + apply (rule continuous_intros | simp add: assms)+ + done + then show ?thesis + apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) + apply (rule integrable_continuous [of 0 "1::real", simplified]) + apply (rule continuous_on_eq [where f = "\x. f(linepath a b x)*(b - a)"]) + apply (auto simp: vector_derivative_linepath_within) + done +qed + +lemma has_field_der_id: "((\x. x\<^sup>2 / 2) has_field_derivative x) (at x)" + by (rule has_derivative_imp_has_field_derivative) + (rule derivative_intros | simp)+ + +lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" + apply (rule contour_integral_unique) + using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] + apply (auto simp: field_simps has_field_der_id) + done + +lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" + by (simp add: contour_integrable_continuous_linepath) + +lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" + by (simp add: contour_integrable_continuous_linepath) + +subsection\<^marker>\tag unimportant\ \Arithmetical combining theorems\ + +lemma has_contour_integral_neg: + "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" + by (simp add: has_integral_neg has_contour_integral_def) + +lemma has_contour_integral_add: + "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ + \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" + by (simp add: has_integral_add has_contour_integral_def algebra_simps) + +lemma has_contour_integral_diff: + "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ + \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" + by (simp add: has_integral_diff has_contour_integral_def algebra_simps) + +lemma has_contour_integral_lmul: + "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" +apply (simp add: has_contour_integral_def) +apply (drule has_integral_mult_right) +apply (simp add: algebra_simps) +done + +lemma has_contour_integral_rmul: + "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" +apply (drule has_contour_integral_lmul) +apply (simp add: mult.commute) +done + +lemma has_contour_integral_div: + "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" + by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) + +lemma has_contour_integral_eq: + "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" +apply (simp add: path_image_def has_contour_integral_def) +by (metis (no_types, lifting) image_eqI has_integral_eq) + +lemma has_contour_integral_bound_linepath: + assumes "(f has_contour_integral i) (linepath a b)" + "0 \ B" "\x. x \ closed_segment a b \ norm(f x) \ B" + shows "norm i \ B * norm(b - a)" +proof - + { fix x::real + assume x: "0 \ x" "x \ 1" + have "norm (f (linepath a b x)) * + norm (vector_derivative (linepath a b) (at x within {0..1})) \ B * norm (b - a)" + by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) + } note * = this + have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" + apply (rule has_integral_bound + [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) + using assms * unfolding has_contour_integral_def + apply (auto simp: norm_mult) + done + then show ?thesis + by (auto simp: content_real) +qed + +(*UNUSED +lemma has_contour_integral_bound_linepath_strong: + fixes a :: real and f :: "complex \ real" + assumes "(f has_contour_integral i) (linepath a b)" + "finite k" + "0 \ B" "\x::real. x \ closed_segment a b - k \ norm(f x) \ B" + shows "norm i \ B*norm(b - a)" +*) + +lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" + unfolding has_contour_integral_linepath + by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) + +lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" + by (simp add: has_contour_integral_def) + +lemma has_contour_integral_is_0: + "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" + by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto + +lemma has_contour_integral_sum: + "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ + \ ((\x. sum (\a. f a x) s) has_contour_integral sum i s) p" + by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) + +subsection\<^marker>\tag unimportant\ \Operations on path integrals\ + +lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" + by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) + +lemma contour_integral_neg: + "f contour_integrable_on g \ contour_integral g (\x. -(f x)) = -(contour_integral g f)" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) + +lemma contour_integral_add: + "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = + contour_integral g f1 + contour_integral g f2" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) + +lemma contour_integral_diff: + "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = + contour_integral g f1 - contour_integral g f2" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) + +lemma contour_integral_lmul: + shows "f contour_integrable_on g + \ contour_integral g (\x. c * f x) = c*contour_integral g f" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) + +lemma contour_integral_rmul: + shows "f contour_integrable_on g + \ contour_integral g (\x. f x * c) = contour_integral g f * c" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) + +lemma contour_integral_div: + shows "f contour_integrable_on g + \ contour_integral g (\x. f x / c) = contour_integral g f / c" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) + +lemma contour_integral_eq: + "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" + apply (simp add: contour_integral_def) + using has_contour_integral_eq + by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) + +lemma contour_integral_eq_0: + "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" + by (simp add: has_contour_integral_is_0 contour_integral_unique) + +lemma contour_integral_bound_linepath: + shows + "\f contour_integrable_on (linepath a b); + 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ + \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" + apply (rule has_contour_integral_bound_linepath [of f]) + apply (auto simp: has_contour_integral_integral) + done + +lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" + by (simp add: contour_integral_unique has_contour_integral_0) + +lemma contour_integral_sum: + "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ + \ contour_integral p (\x. sum (\a. f a x) s) = sum (\a. contour_integral p (f a)) s" + by (auto simp: contour_integral_unique has_contour_integral_sum has_contour_integral_integral) + +lemma contour_integrable_eq: + "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" + unfolding contour_integrable_on_def + by (metis has_contour_integral_eq) + + +subsection\<^marker>\tag unimportant\ \Arithmetic theorems for path integrability\ + +lemma contour_integrable_neg: + "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" + using has_contour_integral_neg contour_integrable_on_def by blast + +lemma contour_integrable_add: + "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" + using has_contour_integral_add contour_integrable_on_def + by fastforce + +lemma contour_integrable_diff: + "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" + using has_contour_integral_diff contour_integrable_on_def + by fastforce + +lemma contour_integrable_lmul: + "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" + using has_contour_integral_lmul contour_integrable_on_def + by fastforce + +lemma contour_integrable_rmul: + "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" + using has_contour_integral_rmul contour_integrable_on_def + by fastforce + +lemma contour_integrable_div: + "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" + using has_contour_integral_div contour_integrable_on_def + by fastforce + +lemma contour_integrable_sum: + "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ + \ (\x. sum (\a. f a x) s) contour_integrable_on p" + unfolding contour_integrable_on_def + by (metis has_contour_integral_sum) + + +subsection\<^marker>\tag unimportant\ \Reversing a path integral\ + +lemma has_contour_integral_reverse_linepath: + "(f has_contour_integral i) (linepath a b) + \ (f has_contour_integral (-i)) (linepath b a)" + using has_contour_integral_reversepath valid_path_linepath by fastforce + +lemma contour_integral_reverse_linepath: + "continuous_on (closed_segment a b) f + \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" +apply (rule contour_integral_unique) +apply (rule has_contour_integral_reverse_linepath) +by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) + + +(* Splitting a path integral in a flat way.*) + +lemma has_contour_integral_split: + assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "(f has_contour_integral (i + j)) (linepath a b)" +proof (cases "k = 0 \ k = 1") + case True + then show ?thesis + using assms by auto +next + case False + then have k: "0 < k" "k < 1" "complex_of_real k \ 1" + using assms by auto + have c': "c = k *\<^sub>R (b - a) + a" + by (metis diff_add_cancel c) + have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" + by (simp add: algebra_simps c') + { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" + have **: "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" + using False apply (simp add: c' algebra_simps) + apply (simp add: real_vector.scale_left_distrib [symmetric] field_split_simps) + done + have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" + using k has_integral_affinity01 [OF *, of "inverse k" "0"] + apply (simp add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) + apply (auto dest: has_integral_cmul [where c = "inverse k"]) + done + } note fi = this + { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" + have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" + using k + apply (simp add: c' field_simps) + apply (simp add: scaleR_conv_of_real divide_simps) + apply (simp add: field_simps) + done + have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" + using k has_integral_affinity01 [OF *, of "inverse(1 - k)" "-(k/(1 - k))"] + apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) + apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) + done + } note fj = this + show ?thesis + using f k + apply (simp add: has_contour_integral_linepath) + apply (simp add: linepath_def) + apply (rule has_integral_combine [OF _ _ fi fj], simp_all) + done +qed + +lemma continuous_on_closed_segment_transform: + assumes f: "continuous_on (closed_segment a b) f" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "continuous_on (closed_segment a c) f" +proof - + have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" + using c by (simp add: algebra_simps) + have "closed_segment a c \ closed_segment a b" + by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) + then show "continuous_on (closed_segment a c) f" + by (rule continuous_on_subset [OF f]) +qed + +lemma contour_integral_split: + assumes f: "continuous_on (closed_segment a b) f" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" +proof - + have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" + using c by (simp add: algebra_simps) + have "closed_segment a c \ closed_segment a b" + by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) + moreover have "closed_segment c b \ closed_segment a b" + by (metis c' ends_in_segment(2) in_segment(1) k subset_closed_segment) + ultimately + have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" + by (auto intro: continuous_on_subset [OF f]) + show ?thesis + by (rule contour_integral_unique) (meson "*" c contour_integrable_continuous_linepath has_contour_integral_integral has_contour_integral_split k) +qed + +lemma contour_integral_split_linepath: + assumes f: "continuous_on (closed_segment a b) f" + and c: "c \ closed_segment a b" + shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" + using c by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) + +subsection\Partial circle path\ + +definition\<^marker>\tag important\ part_circlepath :: "[complex, real, real, real, real] \ complex" + where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" + +lemma pathstart_part_circlepath [simp]: + "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" +by (metis part_circlepath_def pathstart_def pathstart_linepath) + +lemma pathfinish_part_circlepath [simp]: + "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" +by (metis part_circlepath_def pathfinish_def pathfinish_linepath) + +lemma reversepath_part_circlepath[simp]: + "reversepath (part_circlepath z r s t) = part_circlepath z r t s" + unfolding part_circlepath_def reversepath_def linepath_def + by (auto simp:algebra_simps) + +lemma has_vector_derivative_part_circlepath [derivative_intros]: + "((part_circlepath z r s t) has_vector_derivative + (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) + (at x within X)" + apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) + apply (rule has_vector_derivative_real_field) + apply (rule derivative_eq_intros | simp)+ + done + +lemma differentiable_part_circlepath: + "part_circlepath c r a b differentiable at x within A" + using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast + +lemma vector_derivative_part_circlepath: + "vector_derivative (part_circlepath z r s t) (at x) = + \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" + using has_vector_derivative_part_circlepath vector_derivative_at by blast + +lemma vector_derivative_part_circlepath01: + "\0 \ x; x \ 1\ + \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = + \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" + using has_vector_derivative_part_circlepath + by (auto simp: vector_derivative_at_within_ivl) + +lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" + apply (simp add: valid_path_def) + apply (rule C1_differentiable_imp_piecewise) + apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath + intro!: continuous_intros) + done + +lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" + by (simp add: valid_path_imp_path) + +proposition path_image_part_circlepath: + assumes "s \ t" + shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" +proof - + { fix z::real + assume "0 \ z" "z \ 1" + with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" + apply (rule_tac x="(1 - z) * s + z * t" in exI) + apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) + apply (rule conjI) + using mult_right_mono apply blast + using affine_ineq by (metis "mult.commute") + } + moreover + { fix z + assume "s \ z" "z \ t" + then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" + apply (rule_tac x="(z - s)/(t - s)" in image_eqI) + apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) + apply (auto simp: field_split_simps) + done + } + ultimately show ?thesis + by (fastforce simp add: path_image_def part_circlepath_def) +qed + +lemma path_image_part_circlepath': + "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" +proof - + have "path_image (part_circlepath z r s t) = + (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" + by (simp add: image_image path_image_def part_circlepath_def) + also have "linepath s t ` {0..1} = closed_segment s t" + by (rule linepath_image_01) + finally show ?thesis by (simp add: cis_conv_exp) +qed + +lemma path_image_part_circlepath_subset: + "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" +by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) + +lemma in_path_image_part_circlepath: + assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" + shows "norm(w - z) = r" +proof - + have "w \ {c. dist z c = r}" + by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) + thus ?thesis + by (simp add: dist_norm norm_minus_commute) +qed + +lemma path_image_part_circlepath_subset': + assumes "r \ 0" + shows "path_image (part_circlepath z r s t) \ sphere z r" +proof (cases "s \ t") + case True + thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp +next + case False + thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms + by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all +qed + +lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" + by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) + +lemma contour_integral_bound_part_circlepath: + assumes "f contour_integrable_on part_circlepath c r a b" + assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" + shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" +proof - + let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * + exp (\ * linepath a b x))" + have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" + proof (rule integral_norm_bound_integral, goal_cases) + case 1 + with assms(1) show ?case + by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) + next + case (3 x) + with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult + by (intro mult_mono) (auto simp: path_image_def) + qed auto + also have "?I = contour_integral (part_circlepath c r a b) f" + by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) + finally show ?thesis by simp +qed + +lemma has_contour_integral_part_circlepath_iff: + assumes "a < b" + shows "(f has_contour_integral I) (part_circlepath c r a b) \ + ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" +proof - + have "(f has_contour_integral I) (part_circlepath c r a b) \ + ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) + (at x within {0..1})) has_integral I) {0..1}" + unfolding has_contour_integral_def .. + also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * + cis (linepath a b x)) has_integral I) {0..1}" + by (intro has_integral_cong, subst vector_derivative_part_circlepath01) + (simp_all add: cis_conv_exp) + also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * + r * \ * exp (\ * linepath (of_real a) (of_real b) x) * + vector_derivative (linepath (of_real a) (of_real b)) + (at x within {0..1})) has_integral I) {0..1}" + by (intro has_integral_cong, subst vector_derivative_linepath_within) + (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) + also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) + (linepath (of_real a) (of_real b))" + by (simp add: has_contour_integral_def) + also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms + by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) + finally show ?thesis . +qed + +lemma contour_integrable_part_circlepath_iff: + assumes "a < b" + shows "f contour_integrable_on (part_circlepath c r a b) \ + (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (auto simp: contour_integrable_on_def integrable_on_def + has_contour_integral_part_circlepath_iff) + +lemma contour_integral_part_circlepath_eq: + assumes "a < b" + shows "contour_integral (part_circlepath c r a b) f = + integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" +proof (cases "f contour_integrable_on part_circlepath c r a b") + case True + hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (simp add: contour_integrable_part_circlepath_iff) + with True show ?thesis + using has_contour_integral_part_circlepath_iff[OF assms] + contour_integral_unique has_integral_integrable_integral by blast +next + case False + hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (simp add: contour_integrable_part_circlepath_iff) + with False show ?thesis + by (simp add: not_integrable_contour_integral not_integrable_integral) +qed + +lemma contour_integral_part_circlepath_reverse: + "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" + by (subst reversepath_part_circlepath [symmetric], subst contour_integral_reversepath) simp_all + +lemma contour_integral_part_circlepath_reverse': + "b < a \ contour_integral (part_circlepath c r a b) f = + -contour_integral (part_circlepath c r b a) f" + by (rule contour_integral_part_circlepath_reverse) + +lemma finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" +proof (cases "w = 0") + case True then show ?thesis by auto +next + case False + have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" + apply (simp add: norm_mult finite_int_iff_bounded_le) + apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) + apply (auto simp: field_split_simps le_floor_iff) + done + have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" + by blast + show ?thesis + apply (subst exp_Ln [OF False, symmetric]) + apply (simp add: exp_eq) + using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) + done +qed + +lemma finite_bounded_log2: + fixes a::complex + assumes "a \ 0" + shows "finite {z. norm z \ b \ exp(a*z) = w}" +proof - + have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" + by (rule finite_imageI [OF finite_bounded_log]) + show ?thesis + by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) +qed + +lemma has_contour_integral_bound_part_circlepath_strong: + assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" + and "finite k" and le: "0 \ B" "0 < r" "s \ t" + and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" + shows "cmod i \ B * r * (t - s)" +proof - + consider "s = t" | "s < t" using \s \ t\ by linarith + then show ?thesis + proof cases + case 1 with fi [unfolded has_contour_integral] + have "i = 0" by (simp add: vector_derivative_part_circlepath) + with assms show ?thesis by simp + next + case 2 + have [simp]: "\r\ = r" using \r > 0\ by linarith + have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" + by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) + have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y + proof - + define w where "w = (y - z)/of_real r / exp(\ * of_real s)" + have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" + apply (rule finite_vimageI [OF finite_bounded_log2]) + using \s < t\ apply (auto simp: inj_of_real) + done + show ?thesis + apply (simp add: part_circlepath_def linepath_def vimage_def) + apply (rule finite_subset [OF _ fin]) + using le + apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) + done + qed + then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" + by (rule finite_finite_vimage_IntI [OF \finite k\]) + have **: "((\x. if (part_circlepath z r s t x) \ k then 0 + else f(part_circlepath z r s t x) * + vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" + by (rule has_integral_spike [OF negligible_finite [OF fin01]]) (use fi has_contour_integral in auto) + have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" + by (auto intro!: B [unfolded path_image_def image_def, simplified]) + show ?thesis + apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) + using assms apply force + apply (simp add: norm_mult vector_derivative_part_circlepath) + using le * "2" \r > 0\ by auto + qed +qed + +lemma has_contour_integral_bound_part_circlepath: + "\(f has_contour_integral i) (part_circlepath z r s t); + 0 \ B; 0 < r; s \ t; + \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ + \ norm i \ B*r*(t - s)" + by (auto intro: has_contour_integral_bound_part_circlepath_strong) + +lemma contour_integrable_continuous_part_circlepath: + "continuous_on (path_image (part_circlepath z r s t)) f + \ f contour_integrable_on (part_circlepath z r s t)" + apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) + apply (rule integrable_continuous_real) + apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) + done + +lemma simple_path_part_circlepath: + "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" +proof (cases "r = 0 \ s = t") + case True + then show ?thesis + unfolding part_circlepath_def simple_path_def + by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ +next + case False then have "r \ 0" "s \ t" by auto + have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" + by (simp add: algebra_simps) + have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 + \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" + by auto + have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ + (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" + by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] + intro: exI [where x = "-n" for n]) + have 1: "\s - t\ \ 2 * pi" + if "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1" + proof (rule ccontr) + assume "\ \s - t\ \ 2 * pi" + then have *: "\n. t - s \ of_int n * \s - t\" + using False that [of "2*pi / \t - s\"] + by (simp add: abs_minus_commute divide_simps) + show False + using * [of 1] * [of "-1"] by auto + qed + have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n + proof - + have "t-s = 2 * (real_of_int n * pi)/x" + using that by (simp add: field_simps) + then show ?thesis by (metis abs_minus_commute) + qed + have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" + by force + show ?thesis using False + apply (simp add: simple_path_def) + apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) + apply (subst abs_away) + apply (auto simp: 1) + apply (rule ccontr) + apply (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) + done +qed + +lemma arc_part_circlepath: + assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" + shows "arc (part_circlepath z r s t)" +proof - + have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" + and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n + proof (rule ccontr) + assume "x \ y" + have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" + by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_div_cancel_left eq) + then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" + by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) + with \x \ y\ have st: "s-t = (of_int n * (pi * 2) / (y-x))" + by (force simp: field_simps) + have "\real_of_int n\ < \y - x\" + using assms \x \ y\ by (simp add: st abs_mult field_simps) + then show False + using assms x y st by (auto dest: of_int_lessD) + qed + show ?thesis + using assms + apply (simp add: arc_def) + apply (simp add: part_circlepath_def inj_on_def exp_eq) + apply (blast intro: *) + done +qed + +subsection\Special case of one complete circle\ + +definition\<^marker>\tag important\ circlepath :: "[complex, real, real] \ complex" + where "circlepath z r \ part_circlepath z r 0 (2*pi)" + +lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" + by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) + +lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" + by (simp add: circlepath_def) + +lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" + by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) + +lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" +proof - + have "z + of_real r * exp (2 * pi * \ * (x + 1/2)) = + z + of_real r * exp (2 * pi * \ * x + pi * \)" + by (simp add: divide_simps) (simp add: algebra_simps) + also have "\ = z - r * exp (2 * pi * \ * x)" + by (simp add: exp_add) + finally show ?thesis + by (simp add: circlepath path_image_def sphere_def dist_norm) +qed + +lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" + using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] + by (simp add: add.commute) + +lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" + using circlepath_add1 [of z r "x-1/2"] + by (simp add: add.commute) + +lemma path_image_circlepath_minus_subset: + "path_image (circlepath z (-r)) \ path_image (circlepath z r)" + apply (simp add: path_image_def image_def circlepath_minus, clarify) + apply (case_tac "xa \ 1/2", force) + apply (force simp: circlepath_add_half)+ + done + +lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" + using path_image_circlepath_minus_subset by fastforce + +lemma has_vector_derivative_circlepath [derivative_intros]: + "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * of_real x))) + (at x within X)" + apply (simp add: circlepath_def scaleR_conv_of_real) + apply (rule derivative_eq_intros) + apply (simp add: algebra_simps) + done + +lemma vector_derivative_circlepath: + "vector_derivative (circlepath z r) (at x) = + 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" +using has_vector_derivative_circlepath vector_derivative_at by blast + +lemma vector_derivative_circlepath01: + "\0 \ x; x \ 1\ + \ vector_derivative (circlepath z r) (at x within {0..1}) = + 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" + using has_vector_derivative_circlepath + by (auto simp: vector_derivative_at_within_ivl) + +lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" + by (simp add: circlepath_def) + +lemma path_circlepath [simp]: "path (circlepath z r)" + by (simp add: valid_path_imp_path) + +lemma path_image_circlepath_nonneg: + assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" +proof - + have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x + proof (cases "x = z") + case True then show ?thesis by force + next + case False + define w where "w = x - z" + then have "w \ 0" by (simp add: False) + have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" + using cis_conv_exp complex_eq_iff by auto + show ?thesis + apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) + apply (simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) + apply (rule_tac x="t / (2*pi)" in image_eqI) + apply (simp add: field_simps \w \ 0\) + using False ** + apply (auto simp: w_def) + done + qed + show ?thesis + unfolding circlepath path_image_def sphere_def dist_norm + by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) +qed + +lemma path_image_circlepath [simp]: + "path_image (circlepath z r) = sphere z \r\" + using path_image_circlepath_minus + by (force simp: path_image_circlepath_nonneg abs_if) + +lemma has_contour_integral_bound_circlepath_strong: + "\(f has_contour_integral i) (circlepath z r); + finite k; 0 \ B; 0 < r; + \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ + \ norm i \ B*(2*pi*r)" + unfolding circlepath_def + by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) + +lemma has_contour_integral_bound_circlepath: + "\(f has_contour_integral i) (circlepath z r); + 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ + \ norm i \ B*(2*pi*r)" + by (auto intro: has_contour_integral_bound_circlepath_strong) + +lemma contour_integrable_continuous_circlepath: + "continuous_on (path_image (circlepath z r)) f + \ f contour_integrable_on (circlepath z r)" + by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) + +lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" + by (simp add: circlepath_def simple_path_part_circlepath) + +lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" + by (simp add: sphere_def dist_norm norm_minus_commute) + +lemma contour_integral_circlepath: + assumes "r > 0" + shows "contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" +proof (rule contour_integral_unique) + show "((\w. 1 / (w - z)) has_contour_integral 2 * complex_of_real pi * \) (circlepath z r)" + unfolding has_contour_integral_def using assms + apply (subst has_integral_cong) + apply (simp add: vector_derivative_circlepath01) + using has_integral_const_real [of _ 0 1] apply (force simp: circlepath) + done +qed + + +subsection\ Uniform convergence of path integral\ + +text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ + +proposition contour_integral_uniform_limit: + assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" + and ul_f: "uniform_limit (path_image \) f l F" + and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" + and \: "valid_path \" + and [simp]: "\ trivial_limit F" + shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" +proof - + have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) + { fix e::real + assume "0 < e" + then have "0 < e / (\B\ + 1)" by simp + then have "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" + using ul_f [unfolded uniform_limit_iff dist_norm] by auto + with ev_fint + obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" + and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" + using eventually_happens [OF eventually_conj] + by (fastforce simp: contour_integrable_on path_image_def) + have Ble: "B * e / (\B\ + 1) \ e" + using \0 \ B\ \0 < e\ by (simp add: field_split_simps) + have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" + proof (intro exI conjI ballI) + show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" + if "x \ {0..1}" for x + apply (rule order_trans [OF _ Ble]) + using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ + apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) + apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le]) + done + qed (rule inta) + } + then show lintg: "l contour_integrable_on \" + unfolding contour_integrable_on by (metis (mono_tags, lifting)integrable_uniform_limit_real) + { fix e::real + define B' where "B' = B + 1" + have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) + assume "0 < e" + then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" + using ul_f [unfolded uniform_limit_iff dist_norm, rule_format, of "e / B' / 2"] B' + by (simp add: field_simps) + have ie: "integral {0..1::real} (\x. e / 2) < e" using \0 < e\ by simp + have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e / 2" + if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t + proof - + have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" + using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto + also have "\ < e" + by (simp add: B' \0 < e\ mult_imp_div_pos_less) + finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . + then show ?thesis + by (simp add: left_diff_distrib [symmetric] norm_mult) + qed + have le_e: "\x. \\xa\{0..1}. 2 * cmod (f x (\ xa) - l (\ xa)) < e / B'; f x contour_integrable_on \\ + \ cmod (integral {0..1} + (\u. f x (\ u) * vector_derivative \ (at u) - l (\ u) * vector_derivative \ (at u))) < e" + apply (rule le_less_trans [OF integral_norm_bound_integral ie]) + apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) + apply (blast intro: *)+ + done + have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" + apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) + apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) + apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric] le_e) + done + } + then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" + by (rule tendstoI) +qed + +corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: + assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" + and "uniform_limit (sphere z r) f l F" + and "\ trivial_limit F" "0 < r" + shows "l contour_integrable_on (circlepath z r)" + "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" + using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) + + +subsection\<^marker>\tag unimportant\ \General stepping result for derivative formulas\ + +lemma Cauchy_next_derivative: + assumes "continuous_on (path_image \) f'" + and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" + and int: "\w. w \ s - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" + and k: "k \ 0" + and "open s" + and \: "valid_path \" + and w: "w \ s - path_image \" + shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" + and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) + (at w)" (is "?thes2") +proof - + have "open (s - path_image \)" using \open s\ closed_valid_path_image \ by blast + then obtain d where "d>0" and d: "ball w d \ s - path_image \" using w + using open_contains_ball by blast + have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" + by (metis norm_of_nat of_nat_Suc) + have cint: "\x. \x \ w; cmod (x - w) < d\ + \ (\z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \" + 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)+ + have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) + contour_integrable_on \" + unfolding eventually_at + apply (rule_tac x=d in exI) + apply (simp add: \d > 0\ dist_norm field_simps cint) + done + have bim_g: "bounded (image f' (path_image \))" + by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) + then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" + by (force simp: bounded_pos path_image_def) + have twom: "\\<^sub>F n in at w. + \x\path_image \. + cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" + if "0 < e" for e + proof - + have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" + if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" + and uw_less: "cmod (u - w) < e * (d/2) ^ (k+2) / (1 + real k)" + for u x + proof - + define ff where [abs_def]: + "ff n w = + (if n = 0 then inverse(x - w)^k + else if n = 1 then k / (x - w)^(Suc k) + else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w + have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" + by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) + have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d/2))" + if "z \ ball w (d/2)" "i \ 1" for i z + proof - + have "z \ path_image \" + using \x \ path_image \\ d that ball_divide_subset_numeral by blast + then have xz[simp]: "x \ z" using \x \ path_image \\ by blast + then have neq: "x * x + z * z \ x * (z * 2)" + by (blast intro: dest!: sum_sqs_eq) + with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto + then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" + by (simp add: algebra_simps) + show ?thesis using \i \ 1\ + apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) + apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ + done + qed + { fix a::real and b::real assume ab: "a > 0" "b > 0" + then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" + by (subst mult_le_cancel_left_pos) + (use \k \ 0\ in \auto simp: divide_simps\) + with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" + by (simp add: field_simps) + } note canc = this + have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d/2) ^ (k + 2)" + if "v \ ball w (d/2)" for v + proof - + have lessd: "\z. cmod (\ z - v) < d/2 \ cmod (w - \ z) < d" + by (metis that norm_minus_commute norm_triangle_half_r dist_norm mem_ball) + have "d/2 \ cmod (x - v)" using d x that + using lessd d x + by (auto simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps) + then have "d \ cmod (x - v) * 2" + by (simp add: field_split_simps) + then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" + using \0 < d\ order_less_imp_le power_mono by blast + have "x \ v" using that + using \x \ path_image \\ ball_divide_subset_numeral d by fastforce + then show ?thesis + using \d > 0\ apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) + using dpow_le apply (simp add: field_split_simps) + done + qed + have ub: "u \ ball w (d/2)" + using uwd by (simp add: dist_commute dist_norm) + have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d/2) ^ k))" + using complex_Taylor [OF _ ff1 ff2 _ ub, of w, simplified] + by (simp add: ff_def \0 < d\) + then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" + by (simp add: field_simps) + then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + / (cmod (u - w) * real k) + \ (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" + using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) + also have "\ < e" + using uw_less \0 < d\ by (simp add: mult_ac divide_simps) + finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) + / cmod ((u - w) * real k) < e" + by (simp add: norm_mult) + have "x \ u" + using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) + show ?thesis + apply (rule le_less_trans [OF _ e]) + using \k \ 0\ \x \ u\ \u \ w\ + apply (simp add: field_simps norm_divide [symmetric]) + done + qed + show ?thesis + unfolding eventually_at + apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) + apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) + done + qed + have 2: "uniform_limit (path_image \) (\n x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) (\x. f' x / (x - w) ^ Suc k) (at w)" + unfolding uniform_limit_iff dist_norm + proof clarify + fix e::real + assume "0 < e" + have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" + if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" + and x: "0 \ x" "x \ 1" + for u x + proof (cases "(f' (\ x)) = 0") + case True then show ?thesis by (simp add: \0 < e\) + next + case False + have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = + cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k))" + by (simp add: field_simps) + also have "\ = cmod (f' (\ x)) * + cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k)" + by (simp add: norm_mult) + also have "\ < cmod (f' (\ x)) * (e/C)" + using False mult_strict_left_mono [OF ec] by force + also have "\ \ e" using C + by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) + finally show ?thesis . + qed + show "\\<^sub>F n in at w. + \x\path_image \. + cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" + using twom [OF divide_pos_pos [OF \0 < e\ \C > 0\]] unfolding path_image_def + by (force intro: * elim: eventually_mono) + qed + show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" + by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto + have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) + \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" + by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto + have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = + (f u - f w) / (u - w) / k" + if "dist u w < d" for u + proof - + have u: "u \ s - path_image \" + by (metis subsetD d dist_commute mem_ball that) + 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) + done + qed + show ?thes2 + apply (simp add: has_field_derivative_iff del: power_Suc) + apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) + apply (simp add: \k \ 0\ **) + done +qed + +lemma Cauchy_next_derivative_circlepath: + assumes contf: "continuous_on (path_image (circlepath z r)) f" + and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" + and k: "k \ 0" + and w: "w \ ball z r" + shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" + (is "?thes2") +proof - + have "r > 0" using w + using ball_eq_empty by fastforce + have wim: "w \ ball z r - path_image (circlepath z r)" + using w by (auto simp: dist_norm) + show ?thes1 ?thes2 + by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; + auto simp: vector_derivative_circlepath norm_mult)+ +qed + + + +end \ No newline at end of file diff --git a/src/HOL/Analysis/Homeomorphism.thy b/src/HOL/Analysis/Homeomorphism.thy --- a/src/HOL/Analysis/Homeomorphism.thy +++ b/src/HOL/Analysis/Homeomorphism.thy @@ -1,2255 +1,2286 @@ (* Title: HOL/Analysis/Homeomorphism.thy Author: LC Paulson (ported from HOL Light) *) section \Homeomorphism Theorems\ theory Homeomorphism imports Homotopy begin lemma homeomorphic_spheres': fixes a ::"'a::euclidean_space" and b ::"'b::euclidean_space" assumes "0 < \" and dimeq: "DIM('a) = DIM('b)" shows "(sphere a \) homeomorphic (sphere b \)" proof - obtain f :: "'a\'b" and g where "linear f" "linear g" and fg: "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" "\x. g(f x) = x" "\y. f(g y) = y" by (blast intro: isomorphisms_UNIV_UNIV [OF dimeq]) then have "continuous_on UNIV f" "continuous_on UNIV g" using linear_continuous_on linear_linear by blast+ then show ?thesis unfolding homeomorphic_minimal apply(rule_tac x="\x. b + f(x - a)" in exI) apply(rule_tac x="\x. a + g(x - b)" in exI) using assms apply (force intro: continuous_intros continuous_on_compose2 [of _ f] continuous_on_compose2 [of _ g] simp: dist_commute dist_norm fg) done qed lemma homeomorphic_spheres_gen: fixes a :: "'a::euclidean_space" and b :: "'b::euclidean_space" assumes "0 < r" "0 < s" "DIM('a::euclidean_space) = DIM('b::euclidean_space)" shows "(sphere a r homeomorphic sphere b s)" apply (rule homeomorphic_trans [OF homeomorphic_spheres homeomorphic_spheres']) using assms apply auto done subsection \Homeomorphism of all convex compact sets with nonempty interior\ proposition fixes S :: "'a::euclidean_space set" assumes "compact S" and 0: "0 \ rel_interior S" and star: "\x. x \ S \ open_segment 0 x \ rel_interior S" shows starlike_compact_projective1_0: "S - rel_interior S homeomorphic sphere 0 1 \ affine hull S" (is "?SMINUS homeomorphic ?SPHER") and starlike_compact_projective2_0: "S homeomorphic cball 0 1 \ affine hull S" (is "S homeomorphic ?CBALL") proof - have starI: "(u *\<^sub>R x) \ rel_interior S" if "x \ S" "0 \ u" "u < 1" for x u proof (cases "x=0 \ u=0") case True with 0 show ?thesis by force next case False with that show ?thesis by (auto simp: in_segment intro: star [THEN subsetD]) qed have "0 \ S" using assms rel_interior_subset by auto define proj where "proj \ \x::'a. x /\<^sub>R norm x" have eqI: "x = y" if "proj x = proj y" "norm x = norm y" for x y using that by (force simp: proj_def) then have iff_eq: "\x y. (proj x = proj y \ norm x = norm y) \ x = y" by blast have projI: "x \ affine hull S \ proj x \ affine hull S" for x by (metis \0 \ S\ affine_hull_span_0 hull_inc span_mul proj_def) have nproj1 [simp]: "x \ 0 \ norm(proj x) = 1" for x by (simp add: proj_def) have proj0_iff [simp]: "proj x = 0 \ x = 0" for x by (simp add: proj_def) have cont_proj: "continuous_on (UNIV - {0}) proj" unfolding proj_def by (rule continuous_intros | force)+ have proj_spherI: "\x. \x \ affine hull S; x \ 0\ \ proj x \ ?SPHER" by (simp add: projI) have "bounded S" "closed S" using \compact S\ compact_eq_bounded_closed by blast+ have inj_on_proj: "inj_on proj (S - rel_interior S)" proof fix x y assume x: "x \ S - rel_interior S" and y: "y \ S - rel_interior S" and eq: "proj x = proj y" then have xynot: "x \ 0" "y \ 0" "x \ S" "y \ S" "x \ rel_interior S" "y \ rel_interior S" using 0 by auto consider "norm x = norm y" | "norm x < norm y" | "norm x > norm y" by linarith then show "x = y" proof cases assume "norm x = norm y" with iff_eq eq show "x = y" by blast next assume *: "norm x < norm y" have "x /\<^sub>R norm x = norm x *\<^sub>R (x /\<^sub>R norm x) /\<^sub>R norm (norm x *\<^sub>R (x /\<^sub>R norm x))" by force then have "proj ((norm x / norm y) *\<^sub>R y) = proj x" by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) then have [simp]: "(norm x / norm y) *\<^sub>R y = x" by (rule eqI) (simp add: \y \ 0\) have no: "0 \ norm x / norm y" "norm x / norm y < 1" using * by (auto simp: field_split_simps) then show "x = y" using starI [OF \y \ S\ no] xynot by auto next assume *: "norm x > norm y" have "y /\<^sub>R norm y = norm y *\<^sub>R (y /\<^sub>R norm y) /\<^sub>R norm (norm y *\<^sub>R (y /\<^sub>R norm y))" by force then have "proj ((norm y / norm x) *\<^sub>R x) = proj y" by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) then have [simp]: "(norm y / norm x) *\<^sub>R x = y" by (rule eqI) (simp add: \x \ 0\) have no: "0 \ norm y / norm x" "norm y / norm x < 1" using * by (auto simp: field_split_simps) then show "x = y" using starI [OF \x \ S\ no] xynot by auto qed qed have "\surf. homeomorphism (S - rel_interior S) ?SPHER proj surf" proof (rule homeomorphism_compact) show "compact (S - rel_interior S)" using \compact S\ compact_rel_boundary by blast show "continuous_on (S - rel_interior S) proj" using 0 by (blast intro: continuous_on_subset [OF cont_proj]) show "proj ` (S - rel_interior S) = ?SPHER" proof show "proj ` (S - rel_interior S) \ ?SPHER" using 0 by (force simp: hull_inc projI intro: nproj1) show "?SPHER \ proj ` (S - rel_interior S)" proof (clarsimp simp: proj_def) fix x assume "x \ affine hull S" and nox: "norm x = 1" then have "x \ 0" by auto obtain d where "0 < d" and dx: "(d *\<^sub>R x) \ rel_frontier S" and ri: "\e. \0 \ e; e < d\ \ (e *\<^sub>R x) \ rel_interior S" using ray_to_rel_frontier [OF \bounded S\ 0] \x \ affine hull S\ \x \ 0\ by auto show "x \ (\x. x /\<^sub>R norm x) ` (S - rel_interior S)" apply (rule_tac x="d *\<^sub>R x" in image_eqI) using \0 < d\ using dx \closed S\ apply (auto simp: rel_frontier_def field_split_simps nox) done qed qed qed (rule inj_on_proj) then obtain surf where surf: "homeomorphism (S - rel_interior S) ?SPHER proj surf" by blast then have cont_surf: "continuous_on (proj ` (S - rel_interior S)) surf" by (auto simp: homeomorphism_def) have surf_nz: "\x. x \ ?SPHER \ surf x \ 0" by (metis "0" DiffE homeomorphism_def imageI surf) have cont_nosp: "continuous_on (?SPHER) (\x. norm x *\<^sub>R ((surf o proj) x))" apply (rule continuous_intros)+ apply (rule continuous_on_subset [OF cont_proj], force) apply (rule continuous_on_subset [OF cont_surf]) apply (force simp: homeomorphism_image1 [OF surf] dest: proj_spherI) done have surfpS: "\x. \norm x = 1; x \ affine hull S\ \ surf (proj x) \ S" by (metis (full_types) DiffE \0 \ S\ homeomorphism_def image_eqI norm_zero proj_spherI real_vector.scale_zero_left scaleR_one surf) have *: "\y. norm y = 1 \ y \ affine hull S \ x = surf (proj y)" if "x \ S" "x \ rel_interior S" for x proof - have "proj x \ ?SPHER" by (metis (full_types) "0" hull_inc proj_spherI that) moreover have "surf (proj x) = x" by (metis Diff_iff homeomorphism_def surf that) ultimately show ?thesis by (metis \\x. x \ ?SPHER \ surf x \ 0\ hull_inc inverse_1 local.proj_def norm_sgn projI scaleR_one sgn_div_norm that(1)) qed have surfp_notin: "\x. \norm x = 1; x \ affine hull S\ \ surf (proj x) \ rel_interior S" by (metis (full_types) DiffE one_neq_zero homeomorphism_def image_eqI norm_zero proj_spherI surf) have no_sp_im: "(\x. norm x *\<^sub>R surf (proj x)) ` (?SPHER) = S - rel_interior S" by (auto simp: surfpS image_def Bex_def surfp_notin *) have inj_spher: "inj_on (\x. norm x *\<^sub>R surf (proj x)) ?SPHER" proof fix x y assume xy: "x \ ?SPHER" "y \ ?SPHER" and eq: " norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" then have "norm x = 1" "norm y = 1" "x \ affine hull S" "y \ affine hull S" using 0 by auto with eq show "x = y" by (simp add: proj_def) (metis surf xy homeomorphism_def) qed have co01: "compact ?SPHER" by (simp add: compact_Int_closed) show "?SMINUS homeomorphic ?SPHER" apply (subst homeomorphic_sym) apply (rule homeomorphic_compact [OF co01 cont_nosp [unfolded o_def] no_sp_im inj_spher]) done have proj_scaleR: "\a x. 0 < a \ proj (a *\<^sub>R x) = proj x" by (simp add: proj_def) have cont_sp0: "continuous_on (affine hull S - {0}) (surf o proj)" apply (rule continuous_on_compose [OF continuous_on_subset [OF cont_proj]], force) apply (rule continuous_on_subset [OF cont_surf]) using homeomorphism_image1 proj_spherI surf by fastforce obtain B where "B>0" and B: "\x. x \ S \ norm x \ B" by (metis compact_imp_bounded \compact S\ bounded_pos_less less_eq_real_def) have cont_nosp: "continuous (at x within ?CBALL) (\x. norm x *\<^sub>R surf (proj x))" if "norm x \ 1" "x \ affine hull S" for x proof (cases "x=0") case True show ?thesis using True apply (simp add: continuous_within) apply (rule lim_null_scaleR_bounded [where B=B]) apply (simp_all add: tendsto_norm_zero eventually_at) apply (rule_tac x=B in exI) using B surfpS proj_def projI apply (auto simp: \B > 0\) done next case False then have "\\<^sub>F x in at x. (x \ affine hull S - {0}) = (x \ affine hull S)" apply (simp add: eventually_at) apply (rule_tac x="norm x" in exI) apply (auto simp: False) done with cont_sp0 have *: "continuous (at x within affine hull S) (\x. surf (proj x))" apply (simp add: continuous_on_eq_continuous_within) apply (drule_tac x=x in bspec, force simp: False that) apply (simp add: continuous_within Lim_transform_within_set) done show ?thesis apply (rule continuous_within_subset [where s = "affine hull S", OF _ Int_lower2]) apply (rule continuous_intros *)+ done qed have cont_nosp2: "continuous_on ?CBALL (\x. norm x *\<^sub>R ((surf o proj) x))" by (simp add: continuous_on_eq_continuous_within cont_nosp) have "norm y *\<^sub>R surf (proj y) \ S" if "y \ cball 0 1" and yaff: "y \ affine hull S" for y proof (cases "y=0") case True then show ?thesis by (simp add: \0 \ S\) next case False then have "norm y *\<^sub>R surf (proj y) = norm y *\<^sub>R surf (proj (y /\<^sub>R norm y))" by (simp add: proj_def) have "norm y \ 1" using that by simp have "surf (proj (y /\<^sub>R norm y)) \ S" apply (rule surfpS) using proj_def projI yaff by (auto simp: False) then have "surf (proj y) \ S" by (simp add: False proj_def) then show "norm y *\<^sub>R surf (proj y) \ S" by (metis dual_order.antisym le_less_linear norm_ge_zero rel_interior_subset scaleR_one starI subset_eq \norm y \ 1\) qed moreover have "x \ (\x. norm x *\<^sub>R surf (proj x)) ` (?CBALL)" if "x \ S" for x proof (cases "x=0") case True with that hull_inc show ?thesis by fastforce next case False then have psp: "proj (surf (proj x)) = proj x" by (metis homeomorphism_def hull_inc proj_spherI surf that) have nxx: "norm x *\<^sub>R proj x = x" by (simp add: False local.proj_def) have affineI: "(1 / norm (surf (proj x))) *\<^sub>R x \ affine hull S" by (metis \0 \ S\ affine_hull_span_0 hull_inc span_clauses(4) that) have sproj_nz: "surf (proj x) \ 0" by (metis False proj0_iff psp) then have "proj x = proj (proj x)" by (metis False nxx proj_scaleR zero_less_norm_iff) moreover have scaleproj: "\a r. r *\<^sub>R proj a = (r / norm a) *\<^sub>R a" by (simp add: divide_inverse local.proj_def) ultimately have "(norm (surf (proj x)) / norm x) *\<^sub>R x \ rel_interior S" by (metis (no_types) sproj_nz divide_self_if hull_inc norm_eq_zero nproj1 projI psp scaleR_one surfp_notin that) then have "(norm (surf (proj x)) / norm x) \ 1" using starI [OF that] by (meson starI [OF that] le_less_linear norm_ge_zero zero_le_divide_iff) then have nole: "norm x \ norm (surf (proj x))" by (simp add: le_divide_eq_1) show ?thesis apply (rule_tac x="inverse(norm(surf (proj x))) *\<^sub>R x" in image_eqI) apply (metis (no_types, hide_lams) mult.commute scaleproj abs_inverse abs_norm_cancel divide_inverse norm_scaleR nxx positive_imp_inverse_positive proj_scaleR psp sproj_nz zero_less_norm_iff) apply (auto simp: field_split_simps nole affineI) done qed ultimately have im_cball: "(\x. norm x *\<^sub>R surf (proj x)) ` ?CBALL = S" by blast have inj_cball: "inj_on (\x. norm x *\<^sub>R surf (proj x)) ?CBALL" proof fix x y assume "x \ ?CBALL" "y \ ?CBALL" and eq: "norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" then have x: "x \ affine hull S" and y: "y \ affine hull S" using 0 by auto show "x = y" proof (cases "x=0 \ y=0") case True then show "x = y" using eq proj_spherI surf_nz x y by force next case False with x y have speq: "surf (proj x) = surf (proj y)" by (metis eq homeomorphism_apply2 proj_scaleR proj_spherI surf zero_less_norm_iff) then have "norm x = norm y" by (metis \x \ affine hull S\ \y \ affine hull S\ eq proj_spherI real_vector.scale_cancel_right surf_nz) moreover have "proj x = proj y" by (metis (no_types) False speq homeomorphism_apply2 proj_spherI surf x y) ultimately show "x = y" using eq eqI by blast qed qed have co01: "compact ?CBALL" by (simp add: compact_Int_closed) show "S homeomorphic ?CBALL" apply (subst homeomorphic_sym) apply (rule homeomorphic_compact [OF co01 cont_nosp2 [unfolded o_def] im_cball inj_cball]) done qed corollary fixes S :: "'a::euclidean_space set" assumes "compact S" and a: "a \ rel_interior S" and star: "\x. x \ S \ open_segment a x \ rel_interior S" shows starlike_compact_projective1: "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" and starlike_compact_projective2: "S homeomorphic cball a 1 \ affine hull S" proof - have 1: "compact ((+) (-a) ` S)" by (meson assms compact_translation) have 2: "0 \ rel_interior ((+) (-a) ` S)" using a rel_interior_translation [of "- a" S] by (simp cong: image_cong_simp) have 3: "open_segment 0 x \ rel_interior ((+) (-a) ` S)" if "x \ ((+) (-a) ` S)" for x proof - have "x+a \ S" using that by auto then have "open_segment a (x+a) \ rel_interior S" by (metis star) then show ?thesis using open_segment_translation [of a 0 x] using rel_interior_translation [of "- a" S] by (fastforce simp add: ac_simps image_iff cong: image_cong_simp) qed have "S - rel_interior S homeomorphic ((+) (-a) ` S) - rel_interior ((+) (-a) ` S)" by (metis rel_interior_translation translation_diff homeomorphic_translation) also have "... homeomorphic sphere 0 1 \ affine hull ((+) (-a) ` S)" by (rule starlike_compact_projective1_0 [OF 1 2 3]) also have "... = (+) (-a) ` (sphere a 1 \ affine hull S)" by (metis affine_hull_translation left_minus sphere_translation translation_Int) also have "... homeomorphic sphere a 1 \ affine hull S" using homeomorphic_translation homeomorphic_sym by blast finally show "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" . have "S homeomorphic ((+) (-a) ` S)" by (metis homeomorphic_translation) also have "... homeomorphic cball 0 1 \ affine hull ((+) (-a) ` S)" by (rule starlike_compact_projective2_0 [OF 1 2 3]) also have "... = (+) (-a) ` (cball a 1 \ affine hull S)" by (metis affine_hull_translation left_minus cball_translation translation_Int) also have "... homeomorphic cball a 1 \ affine hull S" using homeomorphic_translation homeomorphic_sym by blast finally show "S homeomorphic cball a 1 \ affine hull S" . qed corollary starlike_compact_projective_special: assumes "compact S" and cb01: "cball (0::'a::euclidean_space) 1 \ S" and scale: "\x u. \x \ S; 0 \ u; u < 1\ \ u *\<^sub>R x \ S - frontier S" shows "S homeomorphic (cball (0::'a::euclidean_space) 1)" proof - have "ball 0 1 \ interior S" using cb01 interior_cball interior_mono by blast then have 0: "0 \ rel_interior S" by (meson centre_in_ball subsetD interior_subset_rel_interior le_numeral_extra(2) not_le) have [simp]: "affine hull S = UNIV" using \ball 0 1 \ interior S\ by (auto intro!: affine_hull_nonempty_interior) have star: "open_segment 0 x \ rel_interior S" if "x \ S" for x proof fix p assume "p \ open_segment 0 x" then obtain u where "x \ 0" and u: "0 \ u" "u < 1" and p: "u *\<^sub>R x = p" by (auto simp: in_segment) then show "p \ rel_interior S" using scale [OF that u] closure_subset frontier_def interior_subset_rel_interior by fastforce qed show ?thesis using starlike_compact_projective2_0 [OF \compact S\ 0 star] by simp qed lemma homeomorphic_convex_lemma: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "compact S" "convex T" "compact T" and affeq: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T) \ S homeomorphic T" proof (cases "rel_interior S = {} \ rel_interior T = {}") case True then show ?thesis by (metis Diff_empty affeq \convex S\ \convex T\ aff_dim_empty homeomorphic_empty rel_interior_eq_empty aff_dim_empty) next case False then obtain a b where a: "a \ rel_interior S" and b: "b \ rel_interior T" by auto have starS: "\x. x \ S \ open_segment a x \ rel_interior S" using rel_interior_closure_convex_segment a \convex S\ closure_subset subsetCE by blast have starT: "\x. x \ T \ open_segment b x \ rel_interior T" using rel_interior_closure_convex_segment b \convex T\ closure_subset subsetCE by blast let ?aS = "(+) (-a) ` S" and ?bT = "(+) (-b) ` T" have 0: "0 \ affine hull ?aS" "0 \ affine hull ?bT" by (metis a b subsetD hull_inc image_eqI left_minus rel_interior_subset)+ have subs: "subspace (span ?aS)" "subspace (span ?bT)" by (rule subspace_span)+ moreover have "dim (span ((+) (- a) ` S)) = dim (span ((+) (- b) ` T))" by (metis 0 aff_dim_translation_eq aff_dim_zero affeq dim_span nat_int) ultimately obtain f g where "linear f" "linear g" and fim: "f ` span ?aS = span ?bT" and gim: "g ` span ?bT = span ?aS" and fno: "\x. x \ span ?aS \ norm(f x) = norm x" and gno: "\x. x \ span ?bT \ norm(g x) = norm x" and gf: "\x. x \ span ?aS \ g(f x) = x" and fg: "\x. x \ span ?bT \ f(g x) = x" by (rule isometries_subspaces) blast have [simp]: "continuous_on A f" for A using \linear f\ linear_conv_bounded_linear linear_continuous_on by blast have [simp]: "continuous_on B g" for B using \linear g\ linear_conv_bounded_linear linear_continuous_on by blast have eqspanS: "affine hull ?aS = span ?aS" by (metis a affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) have eqspanT: "affine hull ?bT = span ?bT" by (metis b affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) have "S homeomorphic cball a 1 \ affine hull S" by (rule starlike_compact_projective2 [OF \compact S\ a starS]) also have "... homeomorphic (+) (-a) ` (cball a 1 \ affine hull S)" by (metis homeomorphic_translation) also have "... = cball 0 1 \ (+) (-a) ` (affine hull S)" by (auto simp: dist_norm) also have "... = cball 0 1 \ span ?aS" using eqspanS affine_hull_translation by blast also have "... homeomorphic cball 0 1 \ span ?bT" proof (rule homeomorphicI [where f=f and g=g]) show fim1: "f ` (cball 0 1 \ span ?aS) = cball 0 1 \ span ?bT" apply (rule subset_antisym) using fim fno apply (force simp:, clarify) by (metis IntI fg gim gno image_eqI mem_cball_0) show "g ` (cball 0 1 \ span ?bT) = cball 0 1 \ span ?aS" apply (rule subset_antisym) using gim gno apply (force simp:, clarify) by (metis IntI fim1 gf image_eqI) qed (auto simp: fg gf) also have "... = cball 0 1 \ (+) (-b) ` (affine hull T)" using eqspanT affine_hull_translation by blast also have "... = (+) (-b) ` (cball b 1 \ affine hull T)" by (auto simp: dist_norm) also have "... homeomorphic (cball b 1 \ affine hull T)" by (metis homeomorphic_translation homeomorphic_sym) also have "... homeomorphic T" by (metis starlike_compact_projective2 [OF \compact T\ b starT] homeomorphic_sym) finally have 1: "S homeomorphic T" . have "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" by (rule starlike_compact_projective1 [OF \compact S\ a starS]) also have "... homeomorphic (+) (-a) ` (sphere a 1 \ affine hull S)" by (metis homeomorphic_translation) also have "... = sphere 0 1 \ (+) (-a) ` (affine hull S)" by (auto simp: dist_norm) also have "... = sphere 0 1 \ span ?aS" using eqspanS affine_hull_translation by blast also have "... homeomorphic sphere 0 1 \ span ?bT" proof (rule homeomorphicI [where f=f and g=g]) show fim1: "f ` (sphere 0 1 \ span ?aS) = sphere 0 1 \ span ?bT" apply (rule subset_antisym) using fim fno apply (force simp:, clarify) by (metis IntI fg gim gno image_eqI mem_sphere_0) show "g ` (sphere 0 1 \ span ?bT) = sphere 0 1 \ span ?aS" apply (rule subset_antisym) using gim gno apply (force simp:, clarify) by (metis IntI fim1 gf image_eqI) qed (auto simp: fg gf) also have "... = sphere 0 1 \ (+) (-b) ` (affine hull T)" using eqspanT affine_hull_translation by blast also have "... = (+) (-b) ` (sphere b 1 \ affine hull T)" by (auto simp: dist_norm) also have "... homeomorphic (sphere b 1 \ affine hull T)" by (metis homeomorphic_translation homeomorphic_sym) also have "... homeomorphic T - rel_interior T" by (metis starlike_compact_projective1 [OF \compact T\ b starT] homeomorphic_sym) finally have 2: "S - rel_interior S homeomorphic T - rel_interior T" . show ?thesis using 1 2 by blast qed lemma homeomorphic_convex_compact_sets: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "compact S" "convex T" "compact T" and affeq: "aff_dim S = aff_dim T" shows "S homeomorphic T" using homeomorphic_convex_lemma [OF assms] assms by (auto simp: rel_frontier_def) lemma homeomorphic_rel_frontiers_convex_bounded_sets: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "bounded S" "convex T" "bounded T" and affeq: "aff_dim S = aff_dim T" shows "rel_frontier S homeomorphic rel_frontier T" using assms homeomorphic_convex_lemma [of "closure S" "closure T"] by (simp add: rel_frontier_def convex_rel_interior_closure) subsection\Homeomorphisms between punctured spheres and affine sets\ text\Including the famous stereoscopic projection of the 3-D sphere to the complex plane\ text\The special case with centre 0 and radius 1\ lemma homeomorphic_punctured_affine_sphere_affine_01: assumes "b \ sphere 0 1" "affine T" "0 \ T" "b \ T" "affine p" and affT: "aff_dim T = aff_dim p + 1" shows "(sphere 0 1 \ T) - {b} homeomorphic p" proof - have [simp]: "norm b = 1" "b\b = 1" using assms by (auto simp: norm_eq_1) have [simp]: "T \ {v. b\v = 0} \ {}" using \0 \ T\ by auto have [simp]: "\ T \ {v. b\v = 0}" using \norm b = 1\ \b \ T\ by auto define f where "f \ \x. 2 *\<^sub>R b + (2 / (1 - b\x)) *\<^sub>R (x - b)" define g where "g \ \y. b + (4 / (norm y ^ 2 + 4)) *\<^sub>R (y - 2 *\<^sub>R b)" have [simp]: "\x. \x \ T; b\x = 0\ \ f (g x) = x" unfolding f_def g_def by (simp add: algebra_simps field_split_simps add_nonneg_eq_0_iff) have no: "(norm (f x))\<^sup>2 = 4 * (1 + b \ x) / (1 - b \ x)" if "norm x = 1" and "b \ x \ 1" for x using that apply (simp flip: dot_square_norm add: norm_eq_1 nonzero_eq_divide_eq) apply (simp add: f_def vector_add_divide_simps inner_simps) apply (use sum_sqs_eq [of 1 "b \ x"] in \auto simp add: field_split_simps inner_commute\) done have [simp]: "\u::real. 8 + u * (u * 8) = u * 16 \ u=1" by algebra have [simp]: "\x. \norm x = 1; b \ x \ 1\ \ g (f x) = x" unfolding g_def no by (auto simp: f_def field_split_simps) have [simp]: "norm (g x) = 1" if "x \ T" and "b \ x = 0" for x using that apply (simp only: g_def) apply (rule power2_eq_imp_eq) apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps) apply (simp add: algebra_simps inner_commute) done have [simp]: "b \ g x \ 1" if "x \ T" and "b \ x = 0" for x using that unfolding g_def apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps add_nonneg_eq_0_iff) apply (auto simp: algebra_simps) done have "subspace T" by (simp add: assms subspace_affine) have [simp]: "\x. \x \ T; b \ x = 0\ \ g x \ T" unfolding g_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) have "f ` {x. norm x = 1 \ b\x \ 1} \ {x. b\x = 0}" unfolding f_def using \norm b = 1\ norm_eq_1 by (force simp: field_simps inner_add_right inner_diff_right) moreover have "f ` T \ T" unfolding f_def using assms \subspace T\ by (auto simp add: inner_add_right inner_diff_right mem_affine_3_minus subspace_mul) moreover have "{x. b\x = 0} \ T \ f ` ({x. norm x = 1 \ b\x \ 1} \ T)" apply clarify apply (rule_tac x = "g x" in image_eqI, auto) done ultimately have imf: "f ` ({x. norm x = 1 \ b\x \ 1} \ T) = {x. b\x = 0} \ T" by blast have no4: "\y. b\y = 0 \ norm ((y\y + 4) *\<^sub>R b + 4 *\<^sub>R (y - 2 *\<^sub>R b)) = y\y + 4" apply (rule power2_eq_imp_eq) apply (simp_all add: dot_square_norm [symmetric]) apply (auto simp: power2_eq_square algebra_simps inner_commute) done have [simp]: "\x. \norm x = 1; b \ x \ 1\ \ b \ f x = 0" by (simp add: f_def algebra_simps field_split_simps) have [simp]: "\x. \x \ T; norm x = 1; b \ x \ 1\ \ f x \ T" unfolding f_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) have "g ` {x. b\x = 0} \ {x. norm x = 1 \ b\x \ 1}" unfolding g_def apply (clarsimp simp: no4 vector_add_divide_simps divide_simps add_nonneg_eq_0_iff dot_square_norm [symmetric]) apply (auto simp: algebra_simps) done moreover have "g ` T \ T" unfolding g_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) moreover have "{x. norm x = 1 \ b\x \ 1} \ T \ g ` ({x. b\x = 0} \ T)" apply clarify apply (rule_tac x = "f x" in image_eqI, auto) done ultimately have img: "g ` ({x. b\x = 0} \ T) = {x. norm x = 1 \ b\x \ 1} \ T" by blast have aff: "affine ({x. b\x = 0} \ T)" by (blast intro: affine_hyperplane assms) have contf: "continuous_on ({x. norm x = 1 \ b\x \ 1} \ T) f" unfolding f_def by (rule continuous_intros | force)+ have contg: "continuous_on ({x. b\x = 0} \ T) g" unfolding g_def by (rule continuous_intros | force simp: add_nonneg_eq_0_iff)+ have "(sphere 0 1 \ T) - {b} = {x. norm x = 1 \ (b\x \ 1)} \ T" using \norm b = 1\ by (auto simp: norm_eq_1) (metis vector_eq \b\b = 1\) also have "... homeomorphic {x. b\x = 0} \ T" by (rule homeomorphicI [OF imf img contf contg]) auto also have "... homeomorphic p" apply (rule homeomorphic_affine_sets [OF aff \affine p\]) apply (simp add: Int_commute aff_dim_affine_Int_hyperplane [OF \affine T\] affT) done finally show ?thesis . qed theorem homeomorphic_punctured_affine_sphere_affine: fixes a :: "'a :: euclidean_space" assumes "0 < r" "b \ sphere a r" "affine T" "a \ T" "b \ T" "affine p" and aff: "aff_dim T = aff_dim p + 1" shows "(sphere a r \ T) - {b} homeomorphic p" proof - have "a \ b" using assms by auto then have inj: "inj (\x::'a. x /\<^sub>R norm (a - b))" by (simp add: inj_on_def) have "((sphere a r \ T) - {b}) homeomorphic (+) (-a) ` ((sphere a r \ T) - {b})" by (rule homeomorphic_translation) also have "... homeomorphic (*\<^sub>R) (inverse r) ` (+) (- a) ` (sphere a r \ T - {b})" by (metis \0 < r\ homeomorphic_scaling inverse_inverse_eq inverse_zero less_irrefl) also have "... = sphere 0 1 \ ((*\<^sub>R) (inverse r) ` (+) (- a) ` T) - {(b - a) /\<^sub>R r}" using assms by (auto simp: dist_norm norm_minus_commute divide_simps) also have "... homeomorphic p" apply (rule homeomorphic_punctured_affine_sphere_affine_01) using assms affine_translation [symmetric, of "- a"] aff_dim_translation_eq [of "- a"] apply (auto simp: dist_norm norm_minus_commute affine_scaling inj) done finally show ?thesis . qed corollary homeomorphic_punctured_sphere_affine: fixes a :: "'a :: euclidean_space" assumes "0 < r" and b: "b \ sphere a r" and "affine T" and affS: "aff_dim T + 1 = DIM('a)" shows "(sphere a r - {b}) homeomorphic T" using homeomorphic_punctured_affine_sphere_affine [of r b a UNIV T] assms by auto corollary homeomorphic_punctured_sphere_hyperplane: fixes a :: "'a :: euclidean_space" assumes "0 < r" and b: "b \ sphere a r" and "c \ 0" shows "(sphere a r - {b}) homeomorphic {x::'a. c \ x = d}" apply (rule homeomorphic_punctured_sphere_affine) using assms apply (auto simp: affine_hyperplane of_nat_diff) done proposition homeomorphic_punctured_sphere_affine_gen: fixes a :: "'a :: euclidean_space" assumes "convex S" "bounded S" and a: "a \ rel_frontier S" and "affine T" and affS: "aff_dim S = aff_dim T + 1" shows "rel_frontier S - {a} homeomorphic T" proof - obtain U :: "'a set" where "affine U" "convex U" and affdS: "aff_dim U = aff_dim S" using choose_affine_subset [OF affine_UNIV aff_dim_geq] by (meson aff_dim_affine_hull affine_affine_hull affine_imp_convex) have "S \ {}" using assms by auto then obtain z where "z \ U" by (metis aff_dim_negative_iff equals0I affdS) then have bne: "ball z 1 \ U \ {}" by force then have [simp]: "aff_dim(ball z 1 \ U) = aff_dim U" using aff_dim_convex_Int_open [OF \convex U\ open_ball] by (fastforce simp add: Int_commute) have "rel_frontier S homeomorphic rel_frontier (ball z 1 \ U)" by (rule homeomorphic_rel_frontiers_convex_bounded_sets) (auto simp: \affine U\ affine_imp_convex convex_Int affdS assms) also have "... = sphere z 1 \ U" using convex_affine_rel_frontier_Int [of "ball z 1" U] by (simp add: \affine U\ bne) finally have "rel_frontier S homeomorphic sphere z 1 \ U" . then obtain h k where him: "h ` rel_frontier S = sphere z 1 \ U" and kim: "k ` (sphere z 1 \ U) = rel_frontier S" and hcon: "continuous_on (rel_frontier S) h" and kcon: "continuous_on (sphere z 1 \ U) k" and kh: "\x. x \ rel_frontier S \ k(h(x)) = x" and hk: "\y. y \ sphere z 1 \ U \ h(k(y)) = y" unfolding homeomorphic_def homeomorphism_def by auto have "rel_frontier S - {a} homeomorphic (sphere z 1 \ U) - {h a}" proof (rule homeomorphicI) show h: "h ` (rel_frontier S - {a}) = sphere z 1 \ U - {h a}" using him a kh by auto metis show "k ` (sphere z 1 \ U - {h a}) = rel_frontier S - {a}" by (force simp: h [symmetric] image_comp o_def kh) qed (auto intro: continuous_on_subset hcon kcon simp: kh hk) also have "... homeomorphic T" by (rule homeomorphic_punctured_affine_sphere_affine) (use a him in \auto simp: affS affdS \affine T\ \affine U\ \z \ U\\) finally show ?thesis . qed text\ When dealing with AR, ANR and ANR later, it's useful to know that every set is homeomorphic to a closed subset of a convex set, and if the set is locally compact we can take the convex set to be the universe.\ proposition homeomorphic_closedin_convex: fixes S :: "'m::euclidean_space set" assumes "aff_dim S < DIM('n)" obtains U and T :: "'n::euclidean_space set" where "convex U" "U \ {}" "closedin (top_of_set U) T" "S homeomorphic T" proof (cases "S = {}") case True then show ?thesis by (rule_tac U=UNIV and T="{}" in that) auto next case False then obtain a where "a \ S" by auto obtain i::'n where i: "i \ Basis" "i \ 0" using SOME_Basis Basis_zero by force have "0 \ affine hull ((+) (- a) ` S)" by (simp add: \a \ S\ hull_inc) then have "dim ((+) (- a) ` S) = aff_dim ((+) (- a) ` S)" by (simp add: aff_dim_zero) also have "... < DIM('n)" by (simp add: aff_dim_translation_eq_subtract assms cong: image_cong_simp) finally have dd: "dim ((+) (- a) ` S) < DIM('n)" by linarith have span: "span {x. i \ x = 0} = {x. i \ x = 0}" using span_eq_iff [symmetric, of "{x. i \ x = 0}"] subspace_hyperplane [of i] by simp have "dim ((+) (- a) ` S) \ dim {x. i \ x = 0}" using dd by (simp add: dim_hyperplane [OF \i \ 0\]) then obtain T where "subspace T" and Tsub: "T \ {x. i \ x = 0}" and dimT: "dim T = dim ((+) (- a) ` S)" by (rule choose_subspace_of_subspace) (simp add: span) have "subspace (span ((+) (- a) ` S))" using subspace_span by blast then obtain h k where "linear h" "linear k" and heq: "h ` span ((+) (- a) ` S) = T" and keq:"k ` T = span ((+) (- a) ` S)" and hinv [simp]: "\x. x \ span ((+) (- a) ` S) \ k(h x) = x" and kinv [simp]: "\x. x \ T \ h(k x) = x" apply (rule isometries_subspaces [OF _ \subspace T\]) apply (auto simp: dimT) done have hcont: "continuous_on A h" and kcont: "continuous_on B k" for A B using \linear h\ \linear k\ linear_continuous_on linear_conv_bounded_linear by blast+ have ihhhh[simp]: "\x. x \ S \ i \ h (x - a) = 0" using Tsub [THEN subsetD] heq span_superset by fastforce have "sphere 0 1 - {i} homeomorphic {x. i \ x = 0}" apply (rule homeomorphic_punctured_sphere_affine) using i apply (auto simp: affine_hyperplane) by (metis DIM_positive Suc_eq_plus1 add.left_neutral diff_add_cancel not_le not_less_eq_eq of_nat_1 of_nat_diff) then obtain f g where fg: "homeomorphism (sphere 0 1 - {i}) {x. i \ x = 0} f g" by (force simp: homeomorphic_def) have "h ` (+) (- a) ` S \ T" using heq span_superset span_linear_image by blast then have "g ` h ` (+) (- a) ` S \ g ` {x. i \ x = 0}" using Tsub by (simp add: image_mono) also have "... \ sphere 0 1 - {i}" by (simp add: fg [unfolded homeomorphism_def]) finally have gh_sub_sph: "(g \ h) ` (+) (- a) ` S \ sphere 0 1 - {i}" by (metis image_comp) then have gh_sub_cb: "(g \ h) ` (+) (- a) ` S \ cball 0 1" by (metis Diff_subset order_trans sphere_cball) have [simp]: "\u. u \ S \ norm (g (h (u - a))) = 1" using gh_sub_sph [THEN subsetD] by (auto simp: o_def) have ghcont: "continuous_on ((\x. x - a) ` S) (\x. g (h x))" apply (rule continuous_on_compose2 [OF homeomorphism_cont2 [OF fg] hcont], force) done have kfcont: "continuous_on ((\x. g (h (x - a))) ` S) (\x. k (f x))" apply (rule continuous_on_compose2 [OF kcont]) using homeomorphism_cont1 [OF fg] gh_sub_sph apply (force intro: continuous_on_subset, blast) done have "S homeomorphic (+) (- a) ` S" by (fact homeomorphic_translation) also have "\ homeomorphic (g \ h) ` (+) (- a) ` S" apply (simp add: homeomorphic_def homeomorphism_def cong: image_cong_simp) apply (rule_tac x="g \ h" in exI) apply (rule_tac x="k \ f" in exI) apply (auto simp: ghcont kfcont span_base homeomorphism_apply2 [OF fg] image_comp cong: image_cong_simp) done finally have Shom: "S homeomorphic (\x. g (h x)) ` (\x. x - a) ` S" by (simp cong: image_cong_simp) show ?thesis apply (rule_tac U = "ball 0 1 \ image (g o h) ((+) (- a) ` S)" and T = "image (g o h) ((+) (- a) ` S)" in that) apply (rule convex_intermediate_ball [of 0 1], force) using gh_sub_cb apply force apply force apply (simp add: closedin_closed) apply (rule_tac x="sphere 0 1" in exI) apply (auto simp: Shom cong: image_cong_simp) done qed subsection\Locally compact sets in an open set\ text\ Locally compact sets are closed in an open set and are homeomorphic to an absolutely closed set if we have one more dimension to play with.\ lemma locally_compact_open_Int_closure: fixes S :: "'a :: metric_space set" assumes "locally compact S" obtains T where "open T" "S = T \ closure S" proof - have "\x\S. \T v u. u = S \ T \ x \ u \ u \ v \ v \ S \ open T \ compact v" by (metis assms locally_compact openin_open) then obtain t v where tv: "\x. x \ S \ v x \ S \ open (t x) \ compact (v x) \ (\u. x \ u \ u \ v x \ u = S \ t x)" by metis then have o: "open (\(t ` S))" by blast have "S = \ (v ` S)" using tv by blast also have "... = \(t ` S) \ closure S" proof show "\(v ` S) \ \(t ` S) \ closure S" apply safe apply (metis Int_iff subsetD UN_iff tv) apply (simp add: closure_def rev_subsetD tv) done have "t x \ closure S \ v x" if "x \ S" for x proof - have "t x \ closure S \ closure (t x \ S)" by (simp add: open_Int_closure_subset that tv) also have "... \ v x" by (metis Int_commute closure_minimal compact_imp_closed that tv) finally show ?thesis . qed then show "\(t ` S) \ closure S \ \(v ` S)" by blast qed finally have e: "S = \(t ` S) \ closure S" . show ?thesis by (rule that [OF o e]) qed lemma locally_compact_closedin_open: fixes S :: "'a :: metric_space set" assumes "locally compact S" obtains T where "open T" "closedin (top_of_set T) S" by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int) lemma locally_compact_homeomorphism_projection_closed: assumes "locally compact S" obtains T and f :: "'a \ 'a :: euclidean_space \ 'b :: euclidean_space" where "closed T" "homeomorphism S T f fst" proof (cases "closed S") case True then show ?thesis apply (rule_tac T = "S \ {0}" and f = "\x. (x, 0)" in that) apply (auto simp: closed_Times homeomorphism_def continuous_intros) done next case False obtain U where "open U" and US: "U \ closure S = S" by (metis locally_compact_open_Int_closure [OF assms]) with False have Ucomp: "-U \ {}" using closure_eq by auto have [simp]: "closure (- U) = -U" by (simp add: \open U\ closed_Compl) define f :: "'a \ 'a \ 'b" where "f \ \x. (x, One /\<^sub>R setdist {x} (- U))" have "continuous_on U (\x. (x, One /\<^sub>R setdist {x} (- U)))" apply (intro continuous_intros continuous_on_setdist) by (simp add: Ucomp setdist_eq_0_sing_1) then have homU: "homeomorphism U (f`U) f fst" by (auto simp: f_def homeomorphism_def image_iff continuous_intros) have cloS: "closedin (top_of_set U) S" by (metis US closed_closure closedin_closed_Int) have cont: "isCont ((\x. setdist {x} (- U)) o fst) z" for z :: "'a \ 'b" by (rule continuous_at_compose continuous_intros continuous_at_setdist)+ have setdist1D: "setdist {a} (- U) *\<^sub>R b = One \ setdist {a} (- U) \ 0" for a::'a and b::'b by force have *: "r *\<^sub>R b = One \ b = (1 / r) *\<^sub>R One" for r and b::'b by (metis One_non_0 nonzero_divide_eq_eq real_vector.scale_eq_0_iff real_vector.scale_scale scaleR_one) have "f ` U = (\z. (setdist {fst z} (- U) *\<^sub>R snd z)) -` {One}" apply (auto simp: f_def setdist_eq_0_sing_1 field_simps Ucomp) apply (rule_tac x=a in image_eqI) apply (auto simp: * setdist_eq_0_sing_1 dest: setdist1D) done then have clfU: "closed (f ` U)" apply (rule ssubst) apply (rule continuous_closed_vimage) apply (auto intro: continuous_intros cont [unfolded o_def]) done have "closed (f ` S)" apply (rule closedin_closed_trans [OF _ clfU]) apply (rule homeomorphism_imp_closed_map [OF homU cloS]) done then show ?thesis apply (rule that) apply (rule homeomorphism_of_subsets [OF homU]) using US apply auto done qed lemma locally_compact_closed_Int_open: fixes S :: "'a :: euclidean_space set" shows "locally compact S \ (\U u. closed U \ open u \ S = U \ u)" by (metis closed_closure closed_imp_locally_compact inf_commute locally_compact_Int locally_compact_open_Int_closure open_imp_locally_compact) lemma lowerdim_embeddings: assumes "DIM('a) < DIM('b)" obtains f :: "'a::euclidean_space*real \ 'b::euclidean_space" and g :: "'b \ 'a*real" and j :: 'b where "linear f" "linear g" "\z. g (f z) = z" "j \ Basis" "\x. f(x,0) \ j = 0" proof - let ?B = "Basis :: ('a*real) set" have b01: "(0,1) \ ?B" by (simp add: Basis_prod_def) have "DIM('a * real) \ DIM('b)" by (simp add: Suc_leI assms) then obtain basf :: "'a*real \ 'b" where sbf: "basf ` ?B \ Basis" and injbf: "inj_on basf Basis" by (metis finite_Basis card_le_inj) define basg:: "'b \ 'a * real" where "basg \ \i. if i \ basf ` Basis then inv_into Basis basf i else (0,1)" have bgf[simp]: "basg (basf i) = i" if "i \ Basis" for i using inv_into_f_f injbf that by (force simp: basg_def) have sbg: "basg ` Basis \ ?B" by (force simp: basg_def injbf b01) define f :: "'a*real \ 'b" where "f \ \u. \j\Basis. (u \ basg j) *\<^sub>R j" define g :: "'b \ 'a*real" where "g \ \z. (\i\Basis. (z \ basf i) *\<^sub>R i)" show ?thesis proof show "linear f" unfolding f_def by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps) show "linear g" unfolding g_def by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps) have *: "(\a \ Basis. a \ basf b * (x \ basg a)) = x \ b" if "b \ Basis" for x b using sbf that by auto show gf: "g (f x) = x" for x apply (rule euclidean_eqI) apply (simp add: f_def g_def inner_sum_left scaleR_sum_left algebra_simps) apply (simp add: Groups_Big.sum_distrib_left [symmetric] *) done show "basf(0,1) \ Basis" using b01 sbf by auto then show "f(x,0) \ basf(0,1) = 0" for x apply (simp add: f_def inner_sum_left) apply (rule comm_monoid_add_class.sum.neutral) using b01 inner_not_same_Basis by fastforce qed qed proposition locally_compact_homeomorphic_closed: fixes S :: "'a::euclidean_space set" assumes "locally compact S" and dimlt: "DIM('a) < DIM('b)" obtains T :: "'b::euclidean_space set" where "closed T" "S homeomorphic T" proof - obtain U:: "('a*real)set" and h where "closed U" and homU: "homeomorphism S U h fst" using locally_compact_homeomorphism_projection_closed assms by metis obtain f :: "'a*real \ 'b" and g :: "'b \ 'a*real" where "linear f" "linear g" and gf [simp]: "\z. g (f z) = z" using lowerdim_embeddings [OF dimlt] by metis then have "inj f" by (metis injI) have gfU: "g ` f ` U = U" by (simp add: image_comp o_def) have "S homeomorphic U" using homU homeomorphic_def by blast also have "... homeomorphic f ` U" apply (rule homeomorphicI [OF refl gfU]) apply (meson \inj f\ \linear f\ homeomorphism_cont2 linear_homeomorphism_image) using \linear g\ linear_continuous_on linear_conv_bounded_linear apply blast apply (auto simp: o_def) done finally show ?thesis apply (rule_tac T = "f ` U" in that) apply (rule closed_injective_linear_image [OF \closed U\ \linear f\ \inj f\], assumption) done qed lemma homeomorphic_convex_compact_lemma: fixes S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "cball 0 1 \ S" shows "S homeomorphic (cball (0::'a) 1)" proof (rule starlike_compact_projective_special[OF assms(2-3)]) fix x u assume "x \ S" and "0 \ u" and "u < (1::real)" have "open (ball (u *\<^sub>R x) (1 - u))" by (rule open_ball) moreover have "u *\<^sub>R x \ ball (u *\<^sub>R x) (1 - u)" unfolding centre_in_ball using \u < 1\ by simp moreover have "ball (u *\<^sub>R x) (1 - u) \ S" proof fix y assume "y \ ball (u *\<^sub>R x) (1 - u)" then have "dist (u *\<^sub>R x) y < 1 - u" unfolding mem_ball . with \u < 1\ have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ cball 0 1" by (simp add: dist_norm inverse_eq_divide norm_minus_commute) with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ S" .. with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \ S" using \x \ S\ \0 \ u\ \u < 1\ [THEN less_imp_le] by (rule convexD_alt) then show "y \ S" using \u < 1\ by simp qed ultimately have "u *\<^sub>R x \ interior S" .. then show "u *\<^sub>R x \ S - frontier S" using frontier_def and interior_subset by auto qed proposition homeomorphic_convex_compact_cball: fixes e :: real and S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "interior S \ {}" and "e > 0" shows "S homeomorphic (cball (b::'a) e)" proof - obtain a where "a \ interior S" using assms(3) by auto then obtain d where "d > 0" and d: "cball a d \ S" unfolding mem_interior_cball by auto let ?d = "inverse d" and ?n = "0::'a" have "cball ?n 1 \ (\x. inverse d *\<^sub>R (x - a)) ` S" apply rule apply (rule_tac x="d *\<^sub>R x + a" in image_eqI) defer apply (rule d[unfolded subset_eq, rule_format]) using \d > 0\ unfolding mem_cball dist_norm apply (auto simp add: mult_right_le_one_le) done then have "(\x. inverse d *\<^sub>R (x - a)) ` S homeomorphic cball ?n 1" using homeomorphic_convex_compact_lemma[of "(\x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` S", OF convex_affinity compact_affinity] using assms(1,2) by (auto simp add: scaleR_right_diff_distrib) then show ?thesis apply (rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]]) apply (rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" S "?d *\<^sub>R -a"]]) using \d>0\ \e>0\ apply (auto simp add: scaleR_right_diff_distrib) done qed corollary homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and T :: "'a set" assumes "convex S" "compact S" "interior S \ {}" and "convex T" "compact T" "interior T \ {}" shows "S homeomorphic T" using assms by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym) lemma homeomorphic_closed_intervals: fixes a :: "'a::euclidean_space" and b and c :: "'a::euclidean_space" and d assumes "box a b \ {}" and "box c d \ {}" shows "(cbox a b) homeomorphic (cbox c d)" apply (rule homeomorphic_convex_compact) using assms apply auto done lemma homeomorphic_closed_intervals_real: fixes a::real and b and c::real and d assumes "aCovering spaces and lifting results for them\ definition\<^marker>\tag important\ covering_space :: "'a::topological_space set \ ('a \ 'b) \ 'b::topological_space set \ bool" where "covering_space c p S \ continuous_on c p \ p ` c = S \ (\x \ S. \T. x \ T \ openin (top_of_set S) T \ (\v. \v = c \ p -` T \ (\u \ v. openin (top_of_set c) u) \ pairwise disjnt v \ (\u \ v. \q. homeomorphism u T p q)))" lemma covering_space_imp_continuous: "covering_space c p S \ continuous_on c p" by (simp add: covering_space_def) lemma covering_space_imp_surjective: "covering_space c p S \ p ` c = S" by (simp add: covering_space_def) lemma homeomorphism_imp_covering_space: "homeomorphism S T f g \ covering_space S f T" apply (simp add: homeomorphism_def covering_space_def, clarify) apply (rule_tac x=T in exI, simp) apply (rule_tac x="{S}" in exI, auto) done lemma covering_space_local_homeomorphism: assumes "covering_space c p S" "x \ c" obtains T u q where "x \ T" "openin (top_of_set c) T" "p x \ u" "openin (top_of_set S) u" "homeomorphism T u p q" using assms apply (simp add: covering_space_def, clarify) apply (drule_tac x="p x" in bspec, force) by (metis IntI UnionE vimage_eq) lemma covering_space_local_homeomorphism_alt: assumes p: "covering_space c p S" and "y \ S" obtains x T U q where "p x = y" "x \ T" "openin (top_of_set c) T" "y \ U" "openin (top_of_set S) U" "homeomorphism T U p q" proof - obtain x where "p x = y" "x \ c" using assms covering_space_imp_surjective by blast show ?thesis apply (rule covering_space_local_homeomorphism [OF p \x \ c\]) using that \p x = y\ by blast qed proposition covering_space_open_map: fixes S :: "'a :: metric_space set" and T :: "'b :: metric_space set" assumes p: "covering_space c p S" and T: "openin (top_of_set c) T" shows "openin (top_of_set S) (p ` T)" proof - have pce: "p ` c = S" and covs: "\x. x \ S \ \X VS. x \ X \ openin (top_of_set S) X \ \VS = c \ p -` X \ (\u \ VS. openin (top_of_set c) u) \ pairwise disjnt VS \ (\u \ VS. \q. homeomorphism u X p q)" using p by (auto simp: covering_space_def) have "T \ c" by (metis openin_euclidean_subtopology_iff T) have "\X. openin (top_of_set S) X \ y \ X \ X \ p ` T" if "y \ p ` T" for y proof - have "y \ S" using \T \ c\ pce that by blast obtain U VS where "y \ U" and U: "openin (top_of_set S) U" and VS: "\VS = c \ p -` U" and openVS: "\V \ VS. openin (top_of_set c) V" and homVS: "\V. V \ VS \ \q. homeomorphism V U p q" using covs [OF \y \ S\] by auto obtain x where "x \ c" "p x \ U" "x \ T" "p x = y" apply simp using T [unfolded openin_euclidean_subtopology_iff] \y \ U\ \y \ p ` T\ by blast with VS obtain V where "x \ V" "V \ VS" by auto then obtain q where q: "homeomorphism V U p q" using homVS by blast then have ptV: "p ` (T \ V) = U \ q -` (T \ V)" using VS \V \ VS\ by (auto simp: homeomorphism_def) have ocv: "openin (top_of_set c) V" by (simp add: \V \ VS\ openVS) have "openin (top_of_set U) (U \ q -` (T \ V))" apply (rule continuous_on_open [THEN iffD1, rule_format]) using homeomorphism_def q apply blast using openin_subtopology_Int_subset [of c] q T unfolding homeomorphism_def by (metis inf.absorb_iff2 Int_commute ocv openin_euclidean_subtopology_iff) then have os: "openin (top_of_set S) (U \ q -` (T \ V))" using openin_trans [of U] by (simp add: Collect_conj_eq U) show ?thesis apply (rule_tac x = "p ` (T \ V)" in exI) apply (rule conjI) apply (simp only: ptV os) using \p x = y\ \x \ V\ \x \ T\ apply blast done qed with openin_subopen show ?thesis by blast qed lemma covering_space_lift_unique_gen: fixes f :: "'a::topological_space \ 'b::topological_space" fixes g1 :: "'a \ 'c::real_normed_vector" assumes cov: "covering_space c p S" and eq: "g1 a = g2 a" and f: "continuous_on T f" "f ` T \ S" and g1: "continuous_on T g1" "g1 ` T \ c" and fg1: "\x. x \ T \ f x = p(g1 x)" and g2: "continuous_on T g2" "g2 ` T \ c" and fg2: "\x. x \ T \ f x = p(g2 x)" and u_compt: "U \ components T" and "a \ U" "x \ U" shows "g1 x = g2 x" proof - have "U \ T" by (rule in_components_subset [OF u_compt]) define G12 where "G12 \ {x \ U. g1 x - g2 x = 0}" have "connected U" by (rule in_components_connected [OF u_compt]) have contu: "continuous_on U g1" "continuous_on U g2" using \U \ T\ continuous_on_subset g1 g2 by blast+ have o12: "openin (top_of_set U) G12" unfolding G12_def proof (subst openin_subopen, clarify) fix z assume z: "z \ U" "g1 z - g2 z = 0" obtain v w q where "g1 z \ v" and ocv: "openin (top_of_set c) v" and "p (g1 z) \ w" and osw: "openin (top_of_set S) w" and hom: "homeomorphism v w p q" apply (rule_tac x = "g1 z" in covering_space_local_homeomorphism [OF cov]) using \U \ T\ \z \ U\ g1(2) apply blast+ done have "g2 z \ v" using \g1 z \ v\ z by auto have gg: "U \ g -` v = U \ g -` (v \ g ` U)" for g by auto have "openin (top_of_set (g1 ` U)) (v \ g1 ` U)" using ocv \U \ T\ g1 by (fastforce simp add: openin_open) then have 1: "openin (top_of_set U) (U \ g1 -` v)" unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) have "openin (top_of_set (g2 ` U)) (v \ g2 ` U)" using ocv \U \ T\ g2 by (fastforce simp add: openin_open) then have 2: "openin (top_of_set U) (U \ g2 -` v)" unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) show "\T. openin (top_of_set U) T \ z \ T \ T \ {z \ U. g1 z - g2 z = 0}" using z apply (rule_tac x = "(U \ g1 -` v) \ (U \ g2 -` v)" in exI) apply (intro conjI) apply (rule openin_Int [OF 1 2]) using \g1 z \ v\ \g2 z \ v\ apply (force simp:, clarify) apply (metis \U \ T\ subsetD eq_iff_diff_eq_0 fg1 fg2 hom homeomorphism_def) done qed have c12: "closedin (top_of_set U) G12" unfolding G12_def by (intro continuous_intros continuous_closedin_preimage_constant contu) have "G12 = {} \ G12 = U" by (intro connected_clopen [THEN iffD1, rule_format] \connected U\ conjI o12 c12) with eq \a \ U\ have "\x. x \ U \ g1 x - g2 x = 0" by (auto simp: G12_def) then show ?thesis using \x \ U\ by force qed proposition covering_space_lift_unique: fixes f :: "'a::topological_space \ 'b::topological_space" fixes g1 :: "'a \ 'c::real_normed_vector" assumes "covering_space c p S" "g1 a = g2 a" "continuous_on T f" "f ` T \ S" "continuous_on T g1" "g1 ` T \ c" "\x. x \ T \ f x = p(g1 x)" "continuous_on T g2" "g2 ` T \ c" "\x. x \ T \ f x = p(g2 x)" "connected T" "a \ T" "x \ T" shows "g1 x = g2 x" using covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv by blast lemma covering_space_locally: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes loc: "locally \ C" and cov: "covering_space C p S" and pim: "\T. \T \ C; \ T\ \ \(p ` T)" shows "locally \ S" proof - have "locally \ (p ` C)" apply (rule locally_open_map_image [OF loc]) using cov covering_space_imp_continuous apply blast using cov covering_space_imp_surjective covering_space_open_map apply blast by (simp add: pim) then show ?thesis using covering_space_imp_surjective [OF cov] by metis qed proposition covering_space_locally_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and pim: "\T. \T \ C; \ T\ \ \(p ` T)" and qim: "\q U. \U \ S; continuous_on U q; \ U\ \ \(q ` U)" shows "locally \ S \ locally \ C" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (rule locallyI) fix V x assume V: "openin (top_of_set C) V" and "x \ V" have "p x \ p ` C" by (metis IntE V \x \ V\ imageI openin_open) then obtain T \ where "p x \ T" and opeT: "openin (top_of_set S) T" and veq: "\\ = C \ p -` T" and ope: "\U\\. openin (top_of_set C) U" and hom: "\U\\. \q. homeomorphism U T p q" using cov unfolding covering_space_def by (blast intro: that) have "x \ \\" using V veq \p x \ T\ \x \ V\ openin_imp_subset by fastforce then obtain U where "x \ U" "U \ \" by blast then obtain q where opeU: "openin (top_of_set C) U" and q: "homeomorphism U T p q" using ope hom by blast with V have "openin (top_of_set C) (U \ V)" by blast then have UV: "openin (top_of_set S) (p ` (U \ V))" using cov covering_space_open_map by blast obtain W W' where opeW: "openin (top_of_set S) W" and "\ W'" "p x \ W" "W \ W'" and W'sub: "W' \ p ` (U \ V)" using locallyE [OF L UV] \x \ U\ \x \ V\ by blast then have "W \ T" by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans) show "\U Z. openin (top_of_set C) U \ \ Z \ x \ U \ U \ Z \ Z \ V" proof (intro exI conjI) have "openin (top_of_set T) W" by (meson opeW opeT openin_imp_subset openin_subset_trans \W \ T\) then have "openin (top_of_set U) (q ` W)" by (meson homeomorphism_imp_open_map homeomorphism_symD q) then show "openin (top_of_set C) (q ` W)" using opeU openin_trans by blast show "\ (q ` W')" by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \\ W'\ continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim) show "x \ q ` W" by (metis \p x \ W\ \x \ U\ homeomorphism_def imageI q) show "q ` W \ q ` W'" using \W \ W'\ by blast have "W' \ p ` V" using W'sub by blast then show "q ` W' \ V" using W'sub homeomorphism_apply1 [OF q] by auto qed qed next assume ?rhs then show ?lhs using cov covering_space_locally pim by blast qed lemma covering_space_locally_compact_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally compact S \ locally compact C" apply (rule covering_space_locally_eq [OF assms]) apply (meson assms compact_continuous_image continuous_on_subset covering_space_imp_continuous) using compact_continuous_image by blast lemma covering_space_locally_connected_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally connected S \ locally connected C" apply (rule covering_space_locally_eq [OF assms]) apply (meson connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) using connected_continuous_image by blast lemma covering_space_locally_path_connected_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally path_connected S \ locally path_connected C" apply (rule covering_space_locally_eq [OF assms]) apply (meson path_connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) using path_connected_continuous_image by blast lemma covering_space_locally_compact: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally compact C" "covering_space C p S" shows "locally compact S" using assms covering_space_locally_compact_eq by blast lemma covering_space_locally_connected: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally connected C" "covering_space C p S" shows "locally connected S" using assms covering_space_locally_connected_eq by blast lemma covering_space_locally_path_connected: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally path_connected C" "covering_space C p S" shows "locally path_connected S" using assms covering_space_locally_path_connected_eq by blast proposition covering_space_lift_homotopy: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and h :: "real \ 'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and conth: "continuous_on ({0..1} \ U) h" and him: "h ` ({0..1} \ U) \ S" and heq: "\y. y \ U \ h (0,y) = p(f y)" and contf: "continuous_on U f" and fim: "f ` U \ C" obtains k where "continuous_on ({0..1} \ U) k" "k ` ({0..1} \ U) \ C" "\y. y \ U \ k(0, y) = f y" "\z. z \ {0..1} \ U \ h z = p(k z)" proof - have "\V k. openin (top_of_set U) V \ y \ V \ continuous_on ({0..1} \ V) k \ k ` ({0..1} \ V) \ C \ (\z \ V. k(0, z) = f z) \ (\z \ {0..1} \ V. h z = p(k z))" if "y \ U" for y proof - obtain UU where UU: "\s. s \ S \ s \ (UU s) \ openin (top_of_set S) (UU s) \ (\\. \\ = C \ p -` UU s \ (\U \ \. openin (top_of_set C) U) \ pairwise disjnt \ \ (\U \ \. \q. homeomorphism U (UU s) p q))" using cov unfolding covering_space_def by (metis (mono_tags)) then have ope: "\s. s \ S \ s \ (UU s) \ openin (top_of_set S) (UU s)" by blast have "\k n i. open k \ open n \ t \ k \ y \ n \ i \ S \ h ` (({0..1} \ k) \ (U \ n)) \ UU i" if "t \ {0..1}" for t proof - have hinS: "h (t, y) \ S" using \y \ U\ him that by blast then have "(t,y) \ ({0..1} \ U) \ h -` UU(h(t, y))" using \y \ U\ \t \ {0..1}\ by (auto simp: ope) moreover have ope_01U: "openin (top_of_set ({0..1} \ U)) (({0..1} \ U) \ h -` UU(h(t, y)))" using hinS ope continuous_on_open_gen [OF him] conth by blast ultimately obtain V W where opeV: "open V" and "t \ {0..1} \ V" "t \ {0..1} \ V" and opeW: "open W" and "y \ U" "y \ W" and VW: "({0..1} \ V) \ (U \ W) \ (({0..1} \ U) \ h -` UU(h(t, y)))" by (rule Times_in_interior_subtopology) (auto simp: openin_open) then show ?thesis using hinS by blast qed then obtain K NN X where K: "\t. t \ {0..1} \ open (K t)" and NN: "\t. t \ {0..1} \ open (NN t)" and inUS: "\t. t \ {0..1} \ t \ K t \ y \ NN t \ X t \ S" and him: "\t. t \ {0..1} \ h ` (({0..1} \ K t) \ (U \ NN t)) \ UU (X t)" by (metis (mono_tags)) obtain \ where "\ \ ((\i. K i \ NN i)) ` {0..1}" "finite \" "{0::real..1} \ {y} \ \\" proof (rule compactE) show "compact ({0::real..1} \ {y})" by (simp add: compact_Times) show "{0..1} \ {y} \ (\i\{0..1}. K i \ NN i)" using K inUS by auto show "\B. B \ (\i. K i \ NN i) ` {0..1} \ open B" using K NN by (auto simp: open_Times) qed blast then obtain tk where "tk \ {0..1}" "finite tk" and tk: "{0::real..1} \ {y} \ (\i \ tk. K i \ NN i)" by (metis (no_types, lifting) finite_subset_image) then have "tk \ {}" by auto define n where "n = \(NN ` tk)" have "y \ n" "open n" using inUS NN \tk \ {0..1}\ \finite tk\ by (auto simp: n_def open_INT subset_iff) obtain \ where "0 < \" and \: "\T. \T \ {0..1}; diameter T < \\ \ \B\K ` tk. T \ B" proof (rule Lebesgue_number_lemma [of "{0..1}" "K ` tk"]) show "K ` tk \ {}" using \tk \ {}\ by auto show "{0..1} \ \(K ` tk)" using tk by auto show "\B. B \ K ` tk \ open B" using \tk \ {0..1}\ K by auto qed auto obtain N::nat where N: "N > 1 / \" using reals_Archimedean2 by blast then have "N > 0" using \0 < \\ order.asym by force have *: "\V k. openin (top_of_set U) V \ y \ V \ continuous_on ({0..of_nat n / N} \ V) k \ k ` ({0..of_nat n / N} \ V) \ C \ (\z\V. k (0, z) = f z) \ (\z\{0..of_nat n / N} \ V. h z = p (k z))" if "n \ N" for n using that proof (induction n) case 0 show ?case apply (rule_tac x=U in exI) apply (rule_tac x="f \ snd" in exI) apply (intro conjI \y \ U\ continuous_intros continuous_on_subset [OF contf]) using fim apply (auto simp: heq) done next case (Suc n) then obtain V k where opeUV: "openin (top_of_set U) V" and "y \ V" and contk: "continuous_on ({0..real n / real N} \ V) k" and kim: "k ` ({0..real n / real N} \ V) \ C" and keq: "\z. z \ V \ k (0, z) = f z" and heq: "\z. z \ {0..real n / real N} \ V \ h z = p (k z)" using Suc_leD by auto have "n \ N" using Suc.prems by auto obtain t where "t \ tk" and t: "{real n / real N .. (1 + real n) / real N} \ K t" proof (rule bexE [OF \]) show "{real n / real N .. (1 + real n) / real N} \ {0..1}" using Suc.prems by (auto simp: field_split_simps) show diameter_less: "diameter {real n / real N .. (1 + real n) / real N} < \" using \0 < \\ N by (auto simp: field_split_simps) qed blast have t01: "t \ {0..1}" using \t \ tk\ \tk \ {0..1}\ by blast obtain \ where \: "\\ = C \ p -` UU (X t)" and opeC: "\U. U \ \ \ openin (top_of_set C) U" and "pairwise disjnt \" and homuu: "\U. U \ \ \ \q. homeomorphism U (UU (X t)) p q" using inUS [OF t01] UU by meson have n_div_N_in: "real n / real N \ {real n / real N .. (1 + real n) / real N}" using N by (auto simp: field_split_simps) with t have nN_in_kkt: "real n / real N \ K t" by blast have "k (real n / real N, y) \ C \ p -` UU (X t)" proof (simp, rule conjI) show "k (real n / real N, y) \ C" using \y \ V\ kim keq by force have "p (k (real n / real N, y)) = h (real n / real N, y)" by (simp add: \y \ V\ heq) also have "... \ h ` (({0..1} \ K t) \ (U \ NN t))" apply (rule imageI) using \y \ V\ t01 \n \ N\ apply (simp add: nN_in_kkt \y \ U\ inUS field_split_simps) done also have "... \ UU (X t)" using him t01 by blast finally show "p (k (real n / real N, y)) \ UU (X t)" . qed with \ have "k (real n / real N, y) \ \\" by blast then obtain W where W: "k (real n / real N, y) \ W" and "W \ \" by blast then obtain p' where opeC': "openin (top_of_set C) W" and hom': "homeomorphism W (UU (X t)) p p'" using homuu opeC by blast then have "W \ C" using openin_imp_subset by blast define W' where "W' = UU(X t)" have opeVW: "openin (top_of_set V) (V \ (k \ Pair (n / N)) -` W)" apply (rule continuous_openin_preimage [OF _ _ opeC']) apply (intro continuous_intros continuous_on_subset [OF contk]) using kim apply (auto simp: \y \ V\ W) done obtain N' where opeUN': "openin (top_of_set U) N'" and "y \ N'" and kimw: "k ` ({(real n / real N)} \ N') \ W" apply (rule_tac N' = "(V \ (k \ Pair (n / N)) -` W)" in that) apply (fastforce simp: \y \ V\ W intro!: openin_trans [OF opeVW opeUV])+ done obtain Q Q' where opeUQ: "openin (top_of_set U) Q" and cloUQ': "closedin (top_of_set U) Q'" and "y \ Q" "Q \ Q'" and Q': "Q' \ (U \ NN(t)) \ N' \ V" proof - obtain VO VX where "open VO" "open VX" and VO: "V = U \ VO" and VX: "N' = U \ VX" using opeUV opeUN' by (auto simp: openin_open) then have "open (NN(t) \ VO \ VX)" using NN t01 by blast then obtain e where "e > 0" and e: "cball y e \ NN(t) \ VO \ VX" by (metis Int_iff \N' = U \ VX\ \V = U \ VO\ \y \ N'\ \y \ V\ inUS open_contains_cball t01) show ?thesis proof show "openin (top_of_set U) (U \ ball y e)" by blast show "closedin (top_of_set U) (U \ cball y e)" using e by (auto simp: closedin_closed) qed (use \y \ U\ \e > 0\ VO VX e in auto) qed then have "y \ Q'" "Q \ (U \ NN(t)) \ N' \ V" by blast+ have neq: "{0..real n / real N} \ {real n / real N..(1 + real n) / real N} = {0..(1 + real n) / real N}" apply (auto simp: field_split_simps) by (metis mult_zero_left of_nat_0_le_iff of_nat_0_less_iff order_trans real_mult_le_cancel_iff1) then have neqQ': "{0..real n / real N} \ Q' \ {real n / real N..(1 + real n) / real N} \ Q' = {0..(1 + real n) / real N} \ Q'" by blast have cont: "continuous_on ({0..(1 + real n) / real N} \ Q') (\x. if x \ {0..real n / real N} \ Q' then k x else (p' \ h) x)" unfolding neqQ' [symmetric] proof (rule continuous_on_cases_local, simp_all add: neqQ' del: comp_apply) show "closedin (top_of_set ({0..(1 + real n) / real N} \ Q')) ({0..real n / real N} \ Q')" apply (simp add: closedin_closed) apply (rule_tac x="{0 .. real n / real N} \ UNIV" in exI) using n_div_N_in apply (auto simp: closed_Times) done show "closedin (top_of_set ({0..(1 + real n) / real N} \ Q')) ({real n / real N..(1 + real n) / real N} \ Q')" apply (simp add: closedin_closed) apply (rule_tac x="{real n / real N .. (1 + real n) / real N} \ UNIV" in exI) apply (auto simp: closed_Times) by (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) show "continuous_on ({0..real n / real N} \ Q') k" apply (rule continuous_on_subset [OF contk]) using Q' by auto have "continuous_on ({real n / real N..(1 + real n) / real N} \ Q') h" proof (rule continuous_on_subset [OF conth]) show "{real n / real N..(1 + real n) / real N} \ Q' \ {0..1} \ U" using \N > 0\ apply auto apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) using Suc.prems order_trans apply fastforce apply (metis IntE cloUQ' closedin_closed) done qed moreover have "continuous_on (h ` ({real n / real N..(1 + real n) / real N} \ Q')) p'" proof (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom']]) have "h ` ({real n / real N..(1 + real n) / real N} \ Q') \ h ` (({0..1} \ K t) \ (U \ NN t))" apply (rule image_mono) using \0 < \\ \N > 0\ Suc.prems apply auto apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) using Suc.prems order_trans apply fastforce using t Q' apply auto done with him show "h ` ({real n / real N..(1 + real n) / real N} \ Q') \ UU (X t)" using t01 by blast qed ultimately show "continuous_on ({real n / real N..(1 + real n) / real N} \ Q') (p' \ h)" by (rule continuous_on_compose) have "k (real n / real N, b) = p' (h (real n / real N, b))" if "b \ Q'" for b proof - have "k (real n / real N, b) \ W" using that Q' kimw by force then have "k (real n / real N, b) = p' (p (k (real n / real N, b)))" by (simp add: homeomorphism_apply1 [OF hom']) then show ?thesis using Q' that by (force simp: heq) qed then show "\x. x \ {real n / real N..(1 + real n) / real N} \ Q' \ x \ {0..real n / real N} \ Q' \ k x = (p' \ h) x" by auto qed have h_in_UU: "h (x, y) \ UU (X t)" if "y \ Q" "\ x \ real n / real N" "0 \ x" "x \ (1 + real n) / real N" for x y proof - have "x \ 1" using Suc.prems that order_trans by force moreover have "x \ K t" by (meson atLeastAtMost_iff le_less not_le subset_eq t that) moreover have "y \ U" using \y \ Q\ opeUQ openin_imp_subset by blast moreover have "y \ NN t" using Q' \Q \ Q'\ \y \ Q\ by auto ultimately have "(x, y) \ (({0..1} \ K t) \ (U \ NN t))" using that by auto then have "h (x, y) \ h ` (({0..1} \ K t) \ (U \ NN t))" by blast also have "... \ UU (X t)" by (metis him t01) finally show ?thesis . qed let ?k = "(\x. if x \ {0..real n / real N} \ Q' then k x else (p' \ h) x)" show ?case proof (intro exI conjI) show "continuous_on ({0..real (Suc n) / real N} \ Q) ?k" apply (rule continuous_on_subset [OF cont]) using \Q \ Q'\ by auto have "\a b. \a \ real n / real N; b \ Q'; 0 \ a\ \ k (a, b) \ C" using kim Q' by force moreover have "\a b. \b \ Q; 0 \ a; a \ (1 + real n) / real N; \ a \ real n / real N\ \ p' (h (a, b)) \ C" apply (rule \W \ C\ [THEN subsetD]) using homeomorphism_image2 [OF hom', symmetric] h_in_UU Q' \Q \ Q'\ \W \ C\ apply auto done ultimately show "?k ` ({0..real (Suc n) / real N} \ Q) \ C" using Q' \Q \ Q'\ by force show "\z\Q. ?k (0, z) = f z" using Q' keq \Q \ Q'\ by auto show "\z \ {0..real (Suc n) / real N} \ Q. h z = p(?k z)" using \Q \ U \ NN t \ N' \ V\ heq apply clarsimp using h_in_UU Q' \Q \ Q'\ apply (auto simp: homeomorphism_apply2 [OF hom', symmetric]) done qed (auto simp: \y \ Q\ opeUQ) qed show ?thesis using*[OF order_refl] N \0 < \\ by (simp add: split: if_split_asm) qed then obtain V fs where opeV: "\y. y \ U \ openin (top_of_set U) (V y)" and V: "\y. y \ U \ y \ V y" and contfs: "\y. y \ U \ continuous_on ({0..1} \ V y) (fs y)" and *: "\y. y \ U \ (fs y) ` ({0..1} \ V y) \ C \ (\z \ V y. fs y (0, z) = f z) \ (\z \ {0..1} \ V y. h z = p(fs y z))" by (metis (mono_tags)) then have VU: "\y. y \ U \ V y \ U" by (meson openin_imp_subset) obtain k where contk: "continuous_on ({0..1} \ U) k" and k: "\x i. \i \ U; x \ {0..1} \ U \ {0..1} \ V i\ \ k x = fs i x" proof (rule pasting_lemma_exists) let ?X = "top_of_set ({0..1::real} \ U)" show "topspace ?X \ (\i\U. {0..1} \ V i)" using V by force show "\i. i \ U \ openin (top_of_set ({0..1} \ U)) ({0..1} \ V i)" by (simp add: Abstract_Topology.openin_Times opeV) show "\i. i \ U \ continuous_map (subtopology (top_of_set ({0..1} \ U)) ({0..1} \ V i)) euclidean (fs i)" using contfs apply simp by (metis continuous_map_iff_continuous continuous_on_subset openin_imp_subset openin_subtopology_self subtopology_subtopology) show "fs i x = fs j x" if "i \ U" "j \ U" and x: "x \ topspace ?X \ {0..1} \ V i \ {0..1} \ V j" for i j x proof - obtain u y where "x = (u, y)" "y \ V i" "y \ V j" "0 \ u" "u \ 1" using x by auto show ?thesis proof (rule covering_space_lift_unique [OF cov, of _ "(0,y)" _ "{0..1} \ {y}" h]) show "fs i (0, y) = fs j (0, y)" using*V by (simp add: \y \ V i\ \y \ V j\ that) show conth_y: "continuous_on ({0..1} \ {y}) h" apply (rule continuous_on_subset [OF conth]) using VU \y \ V j\ that by auto show "h ` ({0..1} \ {y}) \ S" using \y \ V i\ assms(3) VU that by fastforce show "continuous_on ({0..1} \ {y}) (fs i)" using continuous_on_subset [OF contfs] \i \ U\ by (simp add: \y \ V i\ subset_iff) show "fs i ` ({0..1} \ {y}) \ C" using "*" \y \ V i\ \i \ U\ by fastforce show "\x. x \ {0..1} \ {y} \ h x = p (fs i x)" using "*" \y \ V i\ \i \ U\ by blast show "continuous_on ({0..1} \ {y}) (fs j)" using continuous_on_subset [OF contfs] \j \ U\ by (simp add: \y \ V j\ subset_iff) show "fs j ` ({0..1} \ {y}) \ C" using "*" \y \ V j\ \j \ U\ by fastforce show "\x. x \ {0..1} \ {y} \ h x = p (fs j x)" using "*" \y \ V j\ \j \ U\ by blast show "connected ({0..1::real} \ {y})" using connected_Icc connected_Times connected_sing by blast show "(0, y) \ {0..1::real} \ {y}" by force show "x \ {0..1} \ {y}" using \x = (u, y)\ x by blast qed qed qed force show ?thesis proof show "k ` ({0..1} \ U) \ C" using V*k VU by fastforce show "\y. y \ U \ k (0, y) = f y" by (simp add: V*k) show "\z. z \ {0..1} \ U \ h z = p (k z)" using V*k by auto qed (auto simp: contk) qed corollary covering_space_lift_homotopy_alt: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and h :: "'c::real_normed_vector \ real \ 'b" assumes cov: "covering_space C p S" and conth: "continuous_on (U \ {0..1}) h" and him: "h ` (U \ {0..1}) \ S" and heq: "\y. y \ U \ h (y,0) = p(f y)" and contf: "continuous_on U f" and fim: "f ` U \ C" obtains k where "continuous_on (U \ {0..1}) k" "k ` (U \ {0..1}) \ C" "\y. y \ U \ k(y, 0) = f y" "\z. z \ U \ {0..1} \ h z = p(k z)" proof - have "continuous_on ({0..1} \ U) (h \ (\z. (snd z, fst z)))" by (intro continuous_intros continuous_on_subset [OF conth]) auto then obtain k where contk: "continuous_on ({0..1} \ U) k" and kim: "k ` ({0..1} \ U) \ C" and k0: "\y. y \ U \ k(0, y) = f y" and heqp: "\z. z \ {0..1} \ U \ (h \ (\z. Pair (snd z) (fst z))) z = p(k z)" apply (rule covering_space_lift_homotopy [OF cov _ _ _ contf fim]) using him by (auto simp: contf heq) show ?thesis apply (rule_tac k="k \ (\z. Pair (snd z) (fst z))" in that) apply (intro continuous_intros continuous_on_subset [OF contk]) using kim heqp apply (auto simp: k0) done qed corollary covering_space_lift_homotopic_function: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and g:: "'c::real_normed_vector \ 'a" assumes cov: "covering_space C p S" and contg: "continuous_on U g" and gim: "g ` U \ C" and pgeq: "\y. y \ U \ p(g y) = f y" and hom: "homotopic_with_canon (\x. True) U S f f'" obtains g' where "continuous_on U g'" "image g' U \ C" "\y. y \ U \ p(g' y) = f' y" proof - obtain h where conth: "continuous_on ({0..1::real} \ U) h" and him: "h ` ({0..1} \ U) \ S" and h0: "\x. h(0, x) = f x" and h1: "\x. h(1, x) = f' x" using hom by (auto simp: homotopic_with_def) have "\y. y \ U \ h (0, y) = p (g y)" by (simp add: h0 pgeq) then obtain k where contk: "continuous_on ({0..1} \ U) k" and kim: "k ` ({0..1} \ U) \ C" and k0: "\y. y \ U \ k(0, y) = g y" and heq: "\z. z \ {0..1} \ U \ h z = p(k z)" using covering_space_lift_homotopy [OF cov conth him _ contg gim] by metis show ?thesis proof show "continuous_on U (k \ Pair 1)" by (meson contk atLeastAtMost_iff continuous_on_o_Pair order_refl zero_le_one) show "(k \ Pair 1) ` U \ C" using kim by auto show "\y. y \ U \ p ((k \ Pair 1) y) = f' y" by (auto simp: h1 heq [symmetric]) qed qed corollary covering_space_lift_inessential_function: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and U :: "'c::real_normed_vector set" assumes cov: "covering_space C p S" and hom: "homotopic_with_canon (\x. True) U S f (\x. a)" obtains g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" proof (cases "U = {}") case True then show ?thesis using that continuous_on_empty by blast next case False then obtain b where b: "b \ C" "p b = a" using covering_space_imp_surjective [OF cov] homotopic_with_imp_subset2 [OF hom] by auto then have gim: "(\y. b) ` U \ C" by blast show ?thesis apply (rule covering_space_lift_homotopic_function [OF cov continuous_on_const gim _ homotopic_with_symD [OF hom]]) using b that apply auto done qed subsection\ Lifting of general functions to covering space\ proposition covering_space_lift_path_strong: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and "a \ C" and "path g" and pag: "path_image g \ S" and pas: "pathstart g = p a" obtains h where "path h" "path_image h \ C" "pathstart h = a" and "\t. t \ {0..1} \ p(h t) = g t" proof - obtain k:: "real \ 'c \ 'a" where contk: "continuous_on ({0..1} \ {undefined}) k" and kim: "k ` ({0..1} \ {undefined}) \ C" and k0: "k (0, undefined) = a" and pk: "\z. z \ {0..1} \ {undefined} \ p(k z) = (g \ fst) z" proof (rule covering_space_lift_homotopy [OF cov, of "{undefined}" "g \ fst"]) show "continuous_on ({0..1::real} \ {undefined::'c}) (g \ fst)" apply (intro continuous_intros) using \path g\ by (simp add: path_def) show "(g \ fst) ` ({0..1} \ {undefined}) \ S" using pag by (auto simp: path_image_def) show "(g \ fst) (0, y) = p a" if "y \ {undefined}" for y::'c by (metis comp_def fst_conv pas pathstart_def) qed (use assms in auto) show ?thesis proof show "path (k \ (\t. Pair t undefined))" unfolding path_def by (intro continuous_on_compose continuous_intros continuous_on_subset [OF contk]) auto show "path_image (k \ (\t. (t, undefined))) \ C" using kim by (auto simp: path_image_def) show "pathstart (k \ (\t. (t, undefined))) = a" by (auto simp: pathstart_def k0) show "\t. t \ {0..1} \ p ((k \ (\t. (t, undefined))) t) = g t" by (auto simp: pk) qed qed corollary covering_space_lift_path: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g" and pig: "path_image g \ S" obtains h where "path h" "path_image h \ C" "\t. t \ {0..1} \ p(h t) = g t" proof - obtain a where "a \ C" "pathstart g = p a" by (metis pig cov covering_space_imp_surjective imageE pathstart_in_path_image subsetCE) show ?thesis using covering_space_lift_path_strong [OF cov \a \ C\ \path g\ pig] by (metis \pathstart g = p a\ that) qed proposition covering_space_lift_homotopic_paths: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g1" and pig1: "path_image g1 \ S" and "path g2" and pig2: "path_image g2 \ S" and hom: "homotopic_paths S g1 g2" and "path h1" and pih1: "path_image h1 \ C" and ph1: "\t. t \ {0..1} \ p(h1 t) = g1 t" and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ p(h2 t) = g2 t" and h1h2: "pathstart h1 = pathstart h2" shows "homotopic_paths C h1 h2" proof - obtain h :: "real \ real \ 'b" where conth: "continuous_on ({0..1} \ {0..1}) h" and him: "h ` ({0..1} \ {0..1}) \ S" and h0: "\x. h (0, x) = g1 x" and h1: "\x. h (1, x) = g2 x" and heq0: "\t. t \ {0..1} \ h (t, 0) = g1 0" and heq1: "\t. t \ {0..1} \ h (t, 1) = g1 1" using hom by (auto simp: homotopic_paths_def homotopic_with_def pathstart_def pathfinish_def) obtain k where contk: "continuous_on ({0..1} \ {0..1}) k" and kim: "k ` ({0..1} \ {0..1}) \ C" and kh2: "\y. y \ {0..1} \ k (y, 0) = h2 0" and hpk: "\z. z \ {0..1} \ {0..1} \ h z = p (k z)" apply (rule covering_space_lift_homotopy_alt [OF cov conth him, of "\x. h2 0"]) using h1h2 ph1 ph2 apply (force simp: heq0 pathstart_def pathfinish_def) using path_image_def pih2 by fastforce+ have contg1: "continuous_on {0..1} g1" and contg2: "continuous_on {0..1} g2" using \path g1\ \path g2\ path_def by blast+ have g1im: "g1 ` {0..1} \ S" and g2im: "g2 ` {0..1} \ S" using path_image_def pig1 pig2 by auto have conth1: "continuous_on {0..1} h1" and conth2: "continuous_on {0..1} h2" using \path h1\ \path h2\ path_def by blast+ have h1im: "h1 ` {0..1} \ C" and h2im: "h2 ` {0..1} \ C" using path_image_def pih1 pih2 by auto show ?thesis unfolding homotopic_paths pathstart_def pathfinish_def proof (intro exI conjI ballI) show keqh1: "k(0, x) = h1 x" if "x \ {0..1}" for x proof (rule covering_space_lift_unique [OF cov _ contg1 g1im]) show "k (0,0) = h1 0" by (metis atLeastAtMost_iff h1h2 kh2 order_refl pathstart_def zero_le_one) show "continuous_on {0..1} (\a. k (0, a))" by (intro continuous_intros continuous_on_compose2 [OF contk]) auto show "\x. x \ {0..1} \ g1 x = p (k (0, x))" by (metis atLeastAtMost_iff h0 hpk zero_le_one mem_Sigma_iff order_refl) qed (use conth1 h1im kim that in \auto simp: ph1\) show "k(1, x) = h2 x" if "x \ {0..1}" for x proof (rule covering_space_lift_unique [OF cov _ contg2 g2im]) show "k (1,0) = h2 0" by (metis atLeastAtMost_iff kh2 order_refl zero_le_one) show "continuous_on {0..1} (\a. k (1, a))" by (intro continuous_intros continuous_on_compose2 [OF contk]) auto show "\x. x \ {0..1} \ g2 x = p (k (1, x))" by (metis atLeastAtMost_iff h1 hpk mem_Sigma_iff order_refl zero_le_one) qed (use conth2 h2im kim that in \auto simp: ph2\) show "\t. t \ {0..1} \ (k \ Pair t) 0 = h1 0" by (metis comp_apply h1h2 kh2 pathstart_def) show "(k \ Pair t) 1 = h1 1" if "t \ {0..1}" for t proof (rule covering_space_lift_unique [OF cov, of "\a. (k \ Pair a) 1" 0 "\a. h1 1" "{0..1}" "\x. g1 1"]) show "(k \ Pair 0) 1 = h1 1" using keqh1 by auto show "continuous_on {0..1} (\a. (k \ Pair a) 1)" apply simp by (intro continuous_intros continuous_on_compose2 [OF contk]) auto show "\x. x \ {0..1} \ g1 1 = p ((k \ Pair x) 1)" using heq1 hpk by auto qed (use contk kim g1im h1im that in \auto simp: ph1\) qed (use contk kim in auto) qed corollary covering_space_monodromy: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g1" and pig1: "path_image g1 \ S" and "path g2" and pig2: "path_image g2 \ S" and hom: "homotopic_paths S g1 g2" and "path h1" and pih1: "path_image h1 \ C" and ph1: "\t. t \ {0..1} \ p(h1 t) = g1 t" and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ p(h2 t) = g2 t" and h1h2: "pathstart h1 = pathstart h2" shows "pathfinish h1 = pathfinish h2" using covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish by blast corollary covering_space_lift_homotopic_path: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and hom: "homotopic_paths S f f'" and "path g" and pig: "path_image g \ C" and a: "pathstart g = a" and b: "pathfinish g = b" and pgeq: "\t. t \ {0..1} \ p(g t) = f t" obtains g' where "path g'" "path_image g' \ C" "pathstart g' = a" "pathfinish g' = b" "\t. t \ {0..1} \ p(g' t) = f' t" proof (rule covering_space_lift_path_strong [OF cov, of a f']) show "a \ C" using a pig by auto show "path f'" "path_image f' \ S" using hom homotopic_paths_imp_path homotopic_paths_imp_subset by blast+ show "pathstart f' = p a" by (metis a atLeastAtMost_iff hom homotopic_paths_imp_pathstart order_refl pathstart_def pgeq zero_le_one) qed (metis (mono_tags, lifting) assms cov covering_space_monodromy hom homotopic_paths_imp_path homotopic_paths_imp_subset pgeq pig) proposition covering_space_lift_general: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and "a \ C" "z \ U" and U: "path_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" and hom: "\r. \path r; path_image r \ U; pathstart r = z; pathfinish r = z\ \ \q. path q \ path_image q \ C \ pathstart q = a \ pathfinish q = a \ homotopic_paths S (f \ r) (p \ q)" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof - have *: "\g h. path g \ path_image g \ U \ pathstart g = z \ pathfinish g = y \ path h \ path_image h \ C \ pathstart h = a \ (\t \ {0..1}. p(h t) = f(g t))" if "y \ U" for y proof - obtain g where "path g" "path_image g \ U" and pastg: "pathstart g = z" and pafig: "pathfinish g = y" using U \z \ U\ \y \ U\ by (force simp: path_connected_def) obtain h where "path h" "path_image h \ C" "pathstart h = a" and "\t. t \ {0..1} \ p(h t) = (f \ g) t" proof (rule covering_space_lift_path_strong [OF cov \a \ C\]) show "path (f \ g)" using \path g\ \path_image g \ U\ contf continuous_on_subset path_continuous_image by blast show "path_image (f \ g) \ S" by (metis \path_image g \ U\ fim image_mono path_image_compose subset_trans) show "pathstart (f \ g) = p a" by (simp add: feq pastg pathstart_compose) qed auto then show ?thesis by (metis \path g\ \path_image g \ U\ comp_apply pafig pastg) qed have "\l. \g h. path g \ path_image g \ U \ pathstart g = z \ pathfinish g = y \ path h \ path_image h \ C \ pathstart h = a \ (\t \ {0..1}. p(h t) = f(g t)) \ pathfinish h = l" for y proof - have "pathfinish h = pathfinish h'" if g: "path g" "path_image g \ U" "pathstart g = z" "pathfinish g = y" and h: "path h" "path_image h \ C" "pathstart h = a" and phg: "\t. t \ {0..1} \ p(h t) = f(g t)" and g': "path g'" "path_image g' \ U" "pathstart g' = z" "pathfinish g' = y" and h': "path h'" "path_image h' \ C" "pathstart h' = a" and phg': "\t. t \ {0..1} \ p(h' t) = f(g' t)" for g h g' h' proof - obtain q where "path q" and piq: "path_image q \ C" and pastq: "pathstart q = a" and pafiq: "pathfinish q = a" and homS: "homotopic_paths S (f \ g +++ reversepath g') (p \ q)" using g g' hom [of "g +++ reversepath g'"] by (auto simp: subset_path_image_join) have papq: "path (p \ q)" using homS homotopic_paths_imp_path by blast have pipq: "path_image (p \ q) \ S" using homS homotopic_paths_imp_subset by blast obtain q' where "path q'" "path_image q' \ C" and "pathstart q' = pathstart q" "pathfinish q' = pathfinish q" and pq'_eq: "\t. t \ {0..1} \ p (q' t) = (f \ g +++ reversepath g') t" using covering_space_lift_homotopic_path [OF cov homotopic_paths_sym [OF homS] \path q\ piq refl refl] by auto have "q' t = (h \ (*\<^sub>R) 2) t" if "0 \ t" "t \ 1/2" for t proof (rule covering_space_lift_unique [OF cov, of q' 0 "h \ (*\<^sub>R) 2" "{0..1/2}" "f \ g \ (*\<^sub>R) 2" t]) show "q' 0 = (h \ (*\<^sub>R) 2) 0" by (metis \pathstart q' = pathstart q\ comp_def h pastq pathstart_def pth_4(2)) show "continuous_on {0..1/2} (f \ g \ (*\<^sub>R) 2)" apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \path g\] continuous_on_subset [OF contf]) using g(2) path_image_def by fastforce+ show "(f \ g \ (*\<^sub>R) 2) ` {0..1/2} \ S" using g(2) path_image_def fim by fastforce show "(h \ (*\<^sub>R) 2) ` {0..1/2} \ C" using h path_image_def by fastforce show "q' ` {0..1/2} \ C" using \path_image q' \ C\ path_image_def by fastforce show "\x. x \ {0..1/2} \ (f \ g \ (*\<^sub>R) 2) x = p (q' x)" by (auto simp: joinpaths_def pq'_eq) show "\x. x \ {0..1/2} \ (f \ g \ (*\<^sub>R) 2) x = p ((h \ (*\<^sub>R) 2) x)" by (simp add: phg) show "continuous_on {0..1/2} q'" by (simp add: continuous_on_path \path q'\) show "continuous_on {0..1/2} (h \ (*\<^sub>R) 2)" apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \path h\], force) done qed (use that in auto) moreover have "q' t = (reversepath h' \ (\t. 2 *\<^sub>R t - 1)) t" if "1/2 < t" "t \ 1" for t proof (rule covering_space_lift_unique [OF cov, of q' 1 "reversepath h' \ (\t. 2 *\<^sub>R t - 1)" "{1/2<..1}" "f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)" t]) show "q' 1 = (reversepath h' \ (\t. 2 *\<^sub>R t - 1)) 1" using h' \pathfinish q' = pathfinish q\ pafiq by (simp add: pathstart_def pathfinish_def reversepath_def) show "continuous_on {1/2<..1} (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1))" apply (intro continuous_intros continuous_on_compose continuous_on_path \path g'\ continuous_on_subset [OF contf]) using g' apply simp_all by (auto simp: path_image_def reversepath_def) show "(f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \ S" using g'(2) path_image_def fim by (auto simp: image_subset_iff path_image_def reversepath_def) show "q' ` {1/2<..1} \ C" using \path_image q' \ C\ path_image_def by fastforce show "(reversepath h' \ (\t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \ C" using h' by (simp add: path_image_def reversepath_def subset_eq) show "\x. x \ {1/2<..1} \ (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) x = p (q' x)" by (auto simp: joinpaths_def pq'_eq) show "\x. x \ {1/2<..1} \ (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) x = p ((reversepath h' \ (\t. 2 *\<^sub>R t - 1)) x)" by (simp add: phg' reversepath_def) show "continuous_on {1/2<..1} q'" by (auto intro: continuous_on_path [OF \path q'\]) show "continuous_on {1/2<..1} (reversepath h' \ (\t. 2 *\<^sub>R t - 1))" apply (intro continuous_intros continuous_on_compose continuous_on_path \path h'\) using h' apply auto done qed (use that in auto) ultimately have "q' t = (h +++ reversepath h') t" if "0 \ t" "t \ 1" for t using that by (simp add: joinpaths_def) then have "path(h +++ reversepath h')" by (auto intro: path_eq [OF \path q'\]) then show ?thesis by (auto simp: \path h\ \path h'\) qed then show ?thesis by metis qed then obtain l :: "'c \ 'a" where l: "\y g h. \path g; path_image g \ U; pathstart g = z; pathfinish g = y; path h; path_image h \ C; pathstart h = a; \t. t \ {0..1} \ p(h t) = f(g t)\ \ pathfinish h = l y" by metis show ?thesis proof show pleq: "p (l y) = f y" if "y \ U" for y using*[OF \y \ U\] by (metis l atLeastAtMost_iff order_refl pathfinish_def zero_le_one) show "l z = a" using l [of "linepath z z" z "linepath a a"] by (auto simp: assms) show LC: "l ` U \ C" by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE) have "\T. openin (top_of_set U) T \ y \ T \ T \ U \ l -` X" if X: "openin (top_of_set C) X" and "y \ U" "l y \ X" for X y proof - have "X \ C" using X openin_euclidean_subtopology_iff by blast have "f y \ S" using fim \y \ U\ by blast then obtain W \ where WV: "f y \ W \ openin (top_of_set S) W \ (\\ = C \ p -` W \ (\U \ \. openin (top_of_set C) U) \ pairwise disjnt \ \ (\U \ \. \q. homeomorphism U W p q))" using cov by (force simp: covering_space_def) then have "l y \ \\" using \X \ C\ pleq that by auto then obtain W' where "l y \ W'" and "W' \ \" by blast with WV obtain p' where opeCW': "openin (top_of_set C) W'" and homUW': "homeomorphism W' W p p'" by blast then have contp': "continuous_on W p'" and p'im: "p' ` W \ W'" using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+ obtain V where "y \ V" "y \ U" and fimW: "f ` V \ W" "V \ U" and "path_connected V" and opeUV: "openin (top_of_set U) V" proof - have "openin (top_of_set U) (U \ f -` W)" using WV contf continuous_on_open_gen fim by auto then show ?thesis using U WV apply (auto simp: locally_path_connected) apply (drule_tac x="U \ f -` W" in spec) apply (drule_tac x=y in spec) apply (auto simp: \y \ U\ intro: that) done qed have "W' \ C" "W \ S" using opeCW' WV openin_imp_subset by auto have p'im: "p' ` W \ W'" using homUW' homeomorphism_image2 by fastforce show ?thesis proof (intro exI conjI) have "openin (top_of_set S) (W \ p' -` (W' \ X))" proof (rule openin_trans) show "openin (top_of_set W) (W \ p' -` (W' \ X))" apply (rule continuous_openin_preimage [OF contp' p'im]) using X \W' \ C\ apply (auto simp: openin_open) done show "openin (top_of_set S) W" using WV by blast qed then show "openin (top_of_set U) (V \ (U \ (f -` (W \ (p' -` (W' \ X))))))" by (blast intro: opeUV openin_subtopology_self continuous_openin_preimage [OF contf fim]) have "p' (f y) \ X" using \l y \ W'\ homeomorphism_apply1 [OF homUW'] pleq \y \ U\ \l y \ X\ by fastforce then show "y \ V \ (U \ f -` (W \ p' -` (W' \ X)))" using \y \ U\ \y \ V\ WV p'im by auto show "V \ (U \ f -` (W \ p' -` (W' \ X))) \ U \ l -` X" proof (intro subsetI IntI; clarify) fix y' assume y': "y' \ V" "y' \ U" "f y' \ W" "p' (f y') \ W'" "p' (f y') \ X" then obtain \ where "path \" "path_image \ \ V" "pathstart \ = y" "pathfinish \ = y'" by (meson \path_connected V\ \y \ V\ path_connected_def) obtain pp qq where "path pp" "path_image pp \ U" "pathstart pp = z" "pathfinish pp = y" "path qq" "path_image qq \ C" "pathstart qq = a" and pqqeq: "\t. t \ {0..1} \ p(qq t) = f(pp t)" using*[OF \y \ U\] by blast have finW: "\x. \0 \ x; x \ 1\ \ f (\ x) \ W" using \path_image \ \ V\ by (auto simp: image_subset_iff path_image_def fimW [THEN subsetD]) have "pathfinish (qq +++ (p' \ f \ \)) = l y'" proof (rule l [of "pp +++ \" y' "qq +++ (p' \ f \ \)"]) show "path (pp +++ \)" by (simp add: \path \\ \path pp\ \pathfinish pp = y\ \pathstart \ = y\) show "path_image (pp +++ \) \ U" using \V \ U\ \path_image \ \ V\ \path_image pp \ U\ not_in_path_image_join by blast show "pathstart (pp +++ \) = z" by (simp add: \pathstart pp = z\) show "pathfinish (pp +++ \) = y'" by (simp add: \pathfinish \ = y'\) have paqq: "pathfinish qq = pathstart (p' \ f \ \)" apply (simp add: \pathstart \ = y\ pathstart_compose) apply (metis (mono_tags, lifting) \l y \ W'\ \path pp\ \path qq\ \path_image pp \ U\ \path_image qq \ C\ \pathfinish pp = y\ \pathstart pp = z\ \pathstart qq = a\ homeomorphism_apply1 [OF homUW'] l pleq pqqeq \y \ U\) done have "continuous_on (path_image \) (p' \ f)" proof (rule continuous_on_compose) show "continuous_on (path_image \) f" using \path_image \ \ V\ \V \ U\ contf continuous_on_subset by blast show "continuous_on (f ` path_image \) p'" apply (rule continuous_on_subset [OF contp']) apply (auto simp: path_image_def pathfinish_def pathstart_def finW) done qed then show "path (qq +++ (p' \ f \ \))" using \path \\ \path qq\ paqq path_continuous_image path_join_imp by blast show "path_image (qq +++ (p' \ f \ \)) \ C" apply (rule subset_path_image_join) apply (simp add: \path_image qq \ C\) by (metis \W' \ C\ \path_image \ \ V\ dual_order.trans fimW(1) image_comp image_mono p'im path_image_compose) show "pathstart (qq +++ (p' \ f \ \)) = a" by (simp add: \pathstart qq = a\) show "p ((qq +++ (p' \ f \ \)) \) = f ((pp +++ \) \)" if \: "\ \ {0..1}" for \ proof (simp add: joinpaths_def, safe) show "p (qq (2*\)) = f (pp (2*\))" if "\*2 \ 1" using \\ \ {0..1}\ pqqeq that by auto show "p (p' (f (\ (2*\ - 1)))) = f (\ (2*\ - 1))" if "\ \*2 \ 1" apply (rule homeomorphism_apply2 [OF homUW' finW]) using that \ by auto qed qed with \pathfinish \ = y'\ \p' (f y') \ X\ show "y' \ l -` X" unfolding pathfinish_join by (simp add: pathfinish_def) qed qed qed then show "continuous_on U l" by (metis IntD1 IntD2 vimage_eq openin_subopen continuous_on_open_gen [OF LC]) qed qed - corollary covering_space_lift_stronger: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" "a \ C" "z \ U" and U: "path_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" and hom: "\r. \path r; path_image r \ U; pathstart r = z; pathfinish r = z\ \ \b. homotopic_paths S (f \ r) (linepath b b)" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof (rule covering_space_lift_general [OF cov U contf fim feq]) fix r assume "path r" "path_image r \ U" "pathstart r = z" "pathfinish r = z" then obtain b where b: "homotopic_paths S (f \ r) (linepath b b)" using hom by blast then have "f (pathstart r) = b" by (metis homotopic_paths_imp_pathstart pathstart_compose pathstart_linepath) then have "homotopic_paths S (f \ r) (linepath (f z) (f z))" by (simp add: b \pathstart r = z\) then have "homotopic_paths S (f \ r) (p \ linepath a a)" by (simp add: o_def feq linepath_def) then show "\q. path q \ path_image q \ C \ pathstart q = a \ pathfinish q = a \ homotopic_paths S (f \ r) (p \ q)" by (force simp: \a \ C\) qed auto corollary covering_space_lift_strong: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" "a \ C" "z \ U" and scU: "simply_connected U" and lpcU: "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq]) show "path_connected U" using scU simply_connected_eq_contractible_loop_some by blast fix r assume r: "path r" "path_image r \ U" "pathstart r = z" "pathfinish r = z" have "linepath (f z) (f z) = f \ linepath z z" by (simp add: o_def linepath_def) then have "homotopic_paths S (f \ r) (linepath (f z) (f z))" by (metis r contf fim homotopic_paths_continuous_image scU simply_connected_eq_contractible_path) then show "\b. homotopic_paths S (f \ r) (linepath b b)" by blast qed blast corollary covering_space_lift: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and U: "simply_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" obtains g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" proof (cases "U = {}") case True with that show ?thesis by auto next case False then obtain z where "z \ U" by blast then obtain a where "a \ C" "f z = p a" by (metis cov covering_space_imp_surjective fim image_iff image_subset_iff) then show ?thesis by (metis that covering_space_lift_strong [OF cov _ \z \ U\ U contf fim]) qed +subsection\<^marker>\tag unimportant\ \Homeomorphisms of arc images\ + +lemma homeomorphism_arc: + fixes g :: "real \ 'a::t2_space" + assumes "arc g" + obtains h where "homeomorphism {0..1} (path_image g) g h" +using assms by (force simp: arc_def homeomorphism_compact path_def path_image_def) + +lemma homeomorphic_arc_image_interval: + fixes g :: "real \ 'a::t2_space" and a::real + assumes "arc g" "a < b" + shows "(path_image g) homeomorphic {a..b}" +proof - + have "(path_image g) homeomorphic {0..1::real}" + by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc) + also have "\ homeomorphic {a..b}" + using assms by (force intro: homeomorphic_closed_intervals_real) + finally show ?thesis . +qed + +lemma homeomorphic_arc_images: + fixes g :: "real \ 'a::t2_space" and h :: "real \ 'b::t2_space" + assumes "arc g" "arc h" + shows "(path_image g) homeomorphic (path_image h)" +proof - + have "(path_image g) homeomorphic {0..1::real}" + by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc) + also have "\ homeomorphic (path_image h)" + by (meson assms homeomorphic_def homeomorphism_arc) + finally show ?thesis . +qed + end diff --git a/src/HOL/Analysis/Path_Connected.thy b/src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy +++ b/src/HOL/Analysis/Path_Connected.thy @@ -1,4006 +1,4007 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Path-Connectedness\ theory Path_Connected imports Starlike T1_Spaces begin subsection \Paths and Arcs\ definition\<^marker>\tag important\ path :: "(real \ 'a::topological_space) \ bool" where "path g \ continuous_on {0..1} g" definition\<^marker>\tag important\ pathstart :: "(real \ 'a::topological_space) \ 'a" where "pathstart g = g 0" definition\<^marker>\tag important\ pathfinish :: "(real \ 'a::topological_space) \ 'a" where "pathfinish g = g 1" definition\<^marker>\tag important\ path_image :: "(real \ 'a::topological_space) \ 'a set" where "path_image g = g ` {0 .. 1}" definition\<^marker>\tag important\ reversepath :: "(real \ 'a::topological_space) \ real \ 'a" where "reversepath g = (\x. g(1 - x))" definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" (infixr "+++" 75) where "g1 +++ g2 = (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" definition\<^marker>\tag important\ simple_path :: "(real \ 'a::topological_space) \ bool" where "simple_path g \ path g \ (\x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" definition\<^marker>\tag important\ arc :: "(real \ 'a :: topological_space) \ bool" where "arc g \ path g \ inj_on g {0..1}" subsection\<^marker>\tag unimportant\\Invariance theorems\ lemma path_eq: "path p \ (\t. t \ {0..1} \ p t = q t) \ path q" using continuous_on_eq path_def by blast lemma path_continuous_image: "path g \ continuous_on (path_image g) f \ path(f \ g)" unfolding path_def path_image_def using continuous_on_compose by blast lemma path_translation_eq: fixes g :: "real \ 'a :: real_normed_vector" shows "path((\x. a + x) \ g) = path g" proof - have g: "g = (\x. -a + x) \ ((\x. a + x) \ g)" by (rule ext) simp show ?thesis unfolding path_def apply safe apply (subst g) apply (rule continuous_on_compose) apply (auto intro: continuous_intros) done qed lemma path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "path(f \ g) = path g" proof - from linear_injective_left_inverse [OF assms] obtain h where h: "linear h" "h \ f = id" by blast then have g: "g = h \ (f \ g)" by (metis comp_assoc id_comp) show ?thesis unfolding path_def using h assms by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear) qed lemma pathstart_translation: "pathstart((\x. a + x) \ g) = a + pathstart g" by (simp add: pathstart_def) lemma pathstart_linear_image_eq: "linear f \ pathstart(f \ g) = f(pathstart g)" by (simp add: pathstart_def) lemma pathfinish_translation: "pathfinish((\x. a + x) \ g) = a + pathfinish g" by (simp add: pathfinish_def) lemma pathfinish_linear_image: "linear f \ pathfinish(f \ g) = f(pathfinish g)" by (simp add: pathfinish_def) lemma path_image_translation: "path_image((\x. a + x) \ g) = (\x. a + x) ` (path_image g)" by (simp add: image_comp path_image_def) lemma path_image_linear_image: "linear f \ path_image(f \ g) = f ` (path_image g)" by (simp add: image_comp path_image_def) lemma reversepath_translation: "reversepath((\x. a + x) \ g) = (\x. a + x) \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma reversepath_linear_image: "linear f \ reversepath(f \ g) = f \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma joinpaths_translation: "((\x. a + x) \ g1) +++ ((\x. a + x) \ g2) = (\x. a + x) \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma joinpaths_linear_image: "linear f \ (f \ g1) +++ (f \ g2) = f \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma simple_path_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "simple_path((\x. a + x) \ g) = simple_path g" by (simp add: simple_path_def path_translation_eq) lemma simple_path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "simple_path(f \ g) = simple_path g" using assms inj_on_eq_iff [of f] by (auto simp: path_linear_image_eq simple_path_def path_translation_eq) lemma arc_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "arc((\x. a + x) \ g) = arc g" by (auto simp: arc_def inj_on_def path_translation_eq) lemma arc_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "arc(f \ g) = arc g" using assms inj_on_eq_iff [of f] by (auto simp: arc_def inj_on_def path_linear_image_eq) subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \ path g" by (simp add: pathin_def path_def) lemma continuous_on_path: "path f \ t \ {0..1} \ continuous_on t f" using continuous_on_subset path_def by blast lemma arc_imp_simple_path: "arc g \ simple_path g" by (simp add: arc_def inj_on_def simple_path_def) lemma arc_imp_path: "arc g \ path g" using arc_def by blast lemma arc_imp_inj_on: "arc g \ inj_on g {0..1}" by (auto simp: arc_def) lemma simple_path_imp_path: "simple_path g \ path g" using simple_path_def by blast lemma simple_path_cases: "simple_path g \ arc g \ pathfinish g = pathstart g" unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def by force lemma simple_path_imp_arc: "simple_path g \ pathfinish g \ pathstart g \ arc g" using simple_path_cases by auto lemma arc_distinct_ends: "arc g \ pathfinish g \ pathstart g" unfolding arc_def inj_on_def pathfinish_def pathstart_def by fastforce lemma arc_simple_path: "arc g \ simple_path g \ pathfinish g \ pathstart g" using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast lemma simple_path_eq_arc: "pathfinish g \ pathstart g \ (simple_path g = arc g)" by (simp add: arc_simple_path) lemma path_image_const [simp]: "path_image (\t. a) = {a}" by (force simp: path_image_def) lemma path_image_nonempty [simp]: "path_image g \ {}" unfolding path_image_def image_is_empty box_eq_empty by auto lemma pathstart_in_path_image[intro]: "pathstart g \ path_image g" unfolding pathstart_def path_image_def by auto lemma pathfinish_in_path_image[intro]: "pathfinish g \ path_image g" unfolding pathfinish_def path_image_def by auto lemma connected_path_image[intro]: "path g \ connected (path_image g)" unfolding path_def path_image_def using connected_continuous_image connected_Icc by blast lemma compact_path_image[intro]: "path g \ compact (path_image g)" unfolding path_def path_image_def using compact_continuous_image connected_Icc by blast lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" unfolding reversepath_def by auto lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g" proof - have *: "\g. path_image (reversepath g) \ path_image g" unfolding path_image_def subset_eq reversepath_def Ball_def image_iff by force show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed lemma path_reversepath [simp]: "path (reversepath g) \ path g" proof - have *: "\g. path g \ path (reversepath g)" unfolding path_def reversepath_def apply (rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"]) done show ?thesis using *[of "reversepath g"] *[of g] unfolding reversepath_reversepath by (rule iffI) qed lemma arc_reversepath: assumes "arc g" shows "arc(reversepath g)" proof - have injg: "inj_on g {0..1}" using assms by (simp add: arc_def) have **: "\x y::real. 1-x = 1-y \ x = y" by simp show ?thesis using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **) qed lemma simple_path_reversepath: "simple_path g \ simple_path (reversepath g)" apply (simp add: simple_path_def) apply (force simp: reversepath_def) done lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" unfolding path_def pathfinish_def pathstart_def proof safe assume cont: "continuous_on {0..1} (g1 +++ g2)" have g1: "continuous_on {0..1} g1 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2))" by (intro continuous_on_cong refl) (auto simp: joinpaths_def) have g2: "continuous_on {0..1} g2 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2 + 1/2))" using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) show "continuous_on {0..1} g1" and "continuous_on {0..1} g2" unfolding g1 g2 by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply) next assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2" have 01: "{0 .. 1} = {0..1/2} \ {1/2 .. 1::real}" by auto { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2) ` {0..1 / 2}" by (intro image_eqI[where x="x/2"]) auto } note 1 = this { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2 - 1) ` {1 / 2..1}" by (intro image_eqI[where x="x/2 + 1/2"]) auto } note 2 = this show "continuous_on {0..1} (g1 +++ g2)" using assms unfolding joinpaths_def 01 apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros) apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) done qed subsection\<^marker>\tag unimportant\ \Path Images\ lemma bounded_path_image: "path g \ bounded(path_image g)" by (simp add: compact_imp_bounded compact_path_image) lemma closed_path_image: fixes g :: "real \ 'a::t2_space" shows "path g \ closed(path_image g)" by (metis compact_path_image compact_imp_closed) lemma connected_simple_path_image: "simple_path g \ connected(path_image g)" by (metis connected_path_image simple_path_imp_path) lemma compact_simple_path_image: "simple_path g \ compact(path_image g)" by (metis compact_path_image simple_path_imp_path) lemma bounded_simple_path_image: "simple_path g \ bounded(path_image g)" by (metis bounded_path_image simple_path_imp_path) lemma closed_simple_path_image: fixes g :: "real \ 'a::t2_space" shows "simple_path g \ closed(path_image g)" by (metis closed_path_image simple_path_imp_path) lemma connected_arc_image: "arc g \ connected(path_image g)" by (metis connected_path_image arc_imp_path) lemma compact_arc_image: "arc g \ compact(path_image g)" by (metis compact_path_image arc_imp_path) lemma bounded_arc_image: "arc g \ bounded(path_image g)" by (metis bounded_path_image arc_imp_path) lemma closed_arc_image: fixes g :: "real \ 'a::t2_space" shows "arc g \ closed(path_image g)" by (metis closed_path_image arc_imp_path) lemma path_image_join_subset: "path_image (g1 +++ g2) \ path_image g1 \ path_image g2" unfolding path_image_def joinpaths_def by auto lemma subset_path_image_join: assumes "path_image g1 \ s" and "path_image g2 \ s" shows "path_image (g1 +++ g2) \ s" using path_image_join_subset[of g1 g2] and assms by auto lemma path_image_join: "pathfinish g1 = pathstart g2 \ path_image(g1 +++ g2) = path_image g1 \ path_image g2" apply (rule subset_antisym [OF path_image_join_subset]) apply (auto simp: pathfinish_def pathstart_def path_image_def joinpaths_def image_def) apply (drule sym) apply (rule_tac x="xa/2" in bexI, auto) apply (rule ccontr) apply (drule_tac x="(xa+1)/2" in bspec) apply (auto simp: field_simps) apply (drule_tac x="1/2" in bspec, auto) done lemma not_in_path_image_join: assumes "x \ path_image g1" and "x \ path_image g2" shows "x \ path_image (g1 +++ g2)" using assms and path_image_join_subset[of g1 g2] by auto lemma pathstart_compose: "pathstart(f \ p) = f(pathstart p)" by (simp add: pathstart_def) lemma pathfinish_compose: "pathfinish(f \ p) = f(pathfinish p)" by (simp add: pathfinish_def) lemma path_image_compose: "path_image (f \ p) = f ` (path_image p)" by (simp add: image_comp path_image_def) lemma path_compose_join: "f \ (p +++ q) = (f \ p) +++ (f \ q)" by (rule ext) (simp add: joinpaths_def) lemma path_compose_reversepath: "f \ reversepath p = reversepath(f \ p)" by (rule ext) (simp add: reversepath_def) lemma joinpaths_eq: "(\t. t \ {0..1} \ p t = p' t) \ (\t. t \ {0..1} \ q t = q' t) \ t \ {0..1} \ (p +++ q) t = (p' +++ q') t" by (auto simp: joinpaths_def) lemma simple_path_inj_on: "simple_path g \ inj_on g {0<..<1}" by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def) subsection\<^marker>\tag unimportant\\Simple paths with the endpoints removed\ lemma simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def Bex_def image_def) apply (metis eq_iff le_less_linear) apply (metis leD linear) using less_eq_real_def zero_le_one apply blast using less_eq_real_def zero_le_one apply blast done lemma connected_simple_path_endless: "simple_path c \ connected(path_image c - {pathstart c,pathfinish c})" apply (simp add: simple_path_endless) apply (rule connected_continuous_image) apply (meson continuous_on_subset greaterThanLessThan_subseteq_atLeastAtMost_iff le_numeral_extra(3) le_numeral_extra(4) path_def simple_path_imp_path) by auto lemma nonempty_simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} \ {}" by (simp add: simple_path_endless) subsection\<^marker>\tag unimportant\\The operations on paths\ lemma path_image_subset_reversepath: "path_image(reversepath g) \ path_image g" by (auto simp: path_image_def reversepath_def) lemma path_imp_reversepath: "path g \ path(reversepath g)" apply (auto simp: path_def reversepath_def) using continuous_on_compose [of "{0..1}" "\x. 1 - x" g] apply (auto simp: continuous_on_op_minus) done lemma half_bounded_equal: "1 \ x * 2 \ x * 2 \ 1 \ x = (1/2::real)" by simp lemma continuous_on_joinpaths: assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2" shows "continuous_on {0..1} (g1 +++ g2)" proof - have *: "{0..1::real} = {0..1/2} \ {1/2..1}" by auto have gg: "g2 0 = g1 1" by (metis assms(3) pathfinish_def pathstart_def) have 1: "continuous_on {0..1/2} (g1 +++ g2)" apply (rule continuous_on_eq [of _ "g1 \ (\x. 2*x)"]) apply (rule continuous_intros | simp add: joinpaths_def assms)+ done have "continuous_on {1/2..1} (g2 \ (\x. 2*x-1))" apply (rule continuous_on_subset [of "{1/2..1}"]) apply (rule continuous_intros | simp add: image_affinity_atLeastAtMost_diff assms)+ done then have 2: "continuous_on {1/2..1} (g1 +++ g2)" apply (rule continuous_on_eq [of "{1/2..1}" "g2 \ (\x. 2*x-1)"]) apply (rule assms continuous_intros | simp add: joinpaths_def mult.commute half_bounded_equal gg)+ done show ?thesis apply (subst *) apply (rule continuous_on_closed_Un) using 1 2 apply auto done qed lemma path_join_imp: "\path g1; path g2; pathfinish g1 = pathstart g2\ \ path(g1 +++ g2)" by (simp) lemma simple_path_join_loop: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" shows "simple_path(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g12: "g1 1 = g2 0" and g21: "g2 1 = g1 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g1 0, g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xyI: "x = 1 \ y \ 0" and xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" "g2 (2 * x - 1) = g1 (2 * y)" have g1im: "g1 (2 * y) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x - 1" in image_eqI, auto) done have False using subsetD [OF sb g1im] xy apply auto apply (drule inj_onD [OF injg1]) using g21 [symmetric] xyI apply (auto dest: inj_onD [OF injg2]) done } note * = this { fix x and y::real assume xy: "y \ 1" "0 \ x" "\ y * 2 \ 1" "x * 2 \ 1" "g1 (2 * x) = g2 (2 * y - 1)" have g1im: "g1 (2 * x) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x" in image_eqI, auto) done have "x = 0 \ y = 1" using subsetD [OF sb g1im] xy apply auto apply (force dest: inj_onD [OF injg1]) using g21 [symmetric] apply (auto dest: inj_onD [OF injg2]) done } note ** = this show ?thesis using assms apply (simp add: arc_def simple_path_def path_join, clarify) apply (simp add: joinpaths_def split: if_split_asm) apply (force dest: inj_onD [OF injg1]) apply (metis *) apply (metis **) apply (force dest: inj_onD [OF injg2]) done qed lemma arc_join: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "path_image g1 \ path_image g2 \ {pathstart g2}" shows "arc(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g11: "g1 1 = g2 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" "g2 (2 * x - 1) = g1 (2 * y)" have g1im: "g1 (2 * y) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x - 1" in image_eqI, auto) done have False using subsetD [OF sb g1im] xy by (auto dest: inj_onD [OF injg2]) } note * = this show ?thesis apply (simp add: arc_def inj_on_def) apply (clarsimp simp add: arc_imp_path assms) apply (simp add: joinpaths_def split: if_split_asm) apply (force dest: inj_onD [OF injg1]) apply (metis *) apply (metis *) apply (force dest: inj_onD [OF injg2]) done qed lemma reversepath_joinpaths: "pathfinish g1 = pathstart g2 \ reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1" unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def by (rule ext) (auto simp: mult.commute) subsection\<^marker>\tag unimportant\\Some reversed and "if and only if" versions of joining theorems\ lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) define e where "e = dist (g1 1) (g2 0)" assume Neg: "pathfinish g1 \ pathstart g2" then have "0 < dist (pathfinish g1) (pathstart g2)" by auto then have "e > 0" by (metis e_def pathfinish_def pathstart_def) then obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; norm x' < d1\ \ dist (g2 x') (g2 0) < e/2" using assms(2) unfolding path_def continuous_on_iff apply (drule_tac x=0 in bspec, simp) by (metis half_gt_zero_iff norm_conv_dist) obtain d2 where "d2 > 0" and d2: "\x'. \x'\{0..1}; dist x' (1/2) < d2\ \ dist ((g1 +++ g2) x') (g1 1) < e/2" using assms(1) \e > 0\ unfolding path_def continuous_on_iff apply (drule_tac x="1/2" in bspec, simp) apply (drule_tac x="e/2" in spec) apply (force simp: joinpaths_def) done have int01_1: "min (1/2) (min d1 d2) / 2 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have [simp]: "\ min (1 / 2) (min d1 d2) \ 0" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def) then have "dist (g1 1) (g2 0) < e/2 + e/2" using dist_triangle_half_r e_def by blast then show False by (simp add: e_def [symmetric]) qed lemma path_join_eq [simp]: fixes g1 :: "real \ 'a::metric_space" assumes "path g1" "path g2" shows "path(g1 +++ g2) \ pathfinish g1 = pathstart g2" using assms by (metis path_join_path_ends path_join_imp) lemma simple_path_joinE: assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2" obtains "arc g1" "arc g2" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" proof - have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have "path g1" using assms path_join simple_path_imp_path by blast moreover have "inj_on g1 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g1 x = g1 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs) qed ultimately have "arc g1" using assms by (simp add: arc_def) have [simp]: "g2 0 = g1 1" using assms by (metis pathfinish_def pathstart_def) have "path g2" using assms path_join simple_path_imp_path by blast moreover have "inj_on g2 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g2 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "(x + 1) / 2" "(y + 1) / 2"] by (force simp: joinpaths_def split_ifs field_split_simps) qed ultimately have "arc g2" using assms by (simp add: arc_def) have "g2 y = g1 0 \ g2 y = g1 1" if "g1 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" for x y using * [of "x / 2" "(y + 1) / 2"] that by (auto simp: joinpaths_def split_ifs field_split_simps) then have "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (fastforce simp: pathstart_def pathfinish_def path_image_def) with \arc g1\ \arc g2\ show ?thesis using that by blast qed lemma simple_path_join_loop_eq: assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" shows "simple_path(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (metis assms simple_path_joinE simple_path_join_loop) lemma arc_join_eq: assumes "pathfinish g1 = pathstart g2" shows "arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g2}" (is "?lhs = ?rhs") proof assume ?lhs then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path) then have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have False if "g1 0 = g2 u" "0 \ u" "u \ 1" for u using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \?lhs\] by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps) then have n1: "pathstart g1 \ path_image g2" unfolding pathstart_def path_image_def using atLeastAtMost_iff by blast show ?rhs using \?lhs\ apply (rule simple_path_joinE [OF arc_imp_simple_path assms]) using n1 by force next assume ?rhs then show ?lhs using assms by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) qed lemma arc_join_eq_alt: "pathfinish g1 = pathstart g2 \ (arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 = {pathstart g2})" using pathfinish_in_path_image by (fastforce simp: arc_join_eq) subsection\<^marker>\tag unimportant\\The joining of paths is associative\ lemma path_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ path(p +++ (q +++ r)) \ path((p +++ q) +++ r)" by simp lemma simple_path_assoc: assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" shows "simple_path (p +++ (q +++ r)) \ simple_path ((p +++ q) +++ r)" proof (cases "pathstart p = pathfinish r") case True show ?thesis proof assume "simple_path (p +++ q +++ r)" with assms True show "simple_path ((p +++ q) +++ r)" by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join dest: arc_distinct_ends [of r]) next assume 0: "simple_path ((p +++ q) +++ r)" with assms True have q: "pathfinish r \ path_image q" using arc_distinct_ends by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join) have "pathstart r \ path_image p" using assms by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff pathfinish_in_path_image pathfinish_join simple_path_joinE) with assms 0 q True show "simple_path (p +++ q +++ r)" by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join dest!: subsetD [OF _ IntI]) qed next case False { fix x :: 'a assume a: "path_image p \ path_image q \ {pathstart q}" "(path_image p \ path_image q) \ path_image r \ {pathstart r}" "x \ path_image p" "x \ path_image r" have "pathstart r \ path_image q" by (metis assms(2) pathfinish_in_path_image) with a have "x = pathstart q" by blast } with False assms show ?thesis by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join) qed lemma arc_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ arc(p +++ (q +++ r)) \ arc((p +++ q) +++ r)" by (simp add: arc_simple_path simple_path_assoc) subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ lemma path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" by auto lemma simple_path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ simple_path(p +++ q) \ simple_path(q +++ p)" by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) lemma path_image_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path_image(p +++ q) = path_image(q +++ p)" by (simp add: path_image_join sup_commute) subsection\Subpath\ definition\<^marker>\tag important\ subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \x. g((b - a) * x + a)" lemma path_image_subpath_gen: fixes g :: "_ \ 'a::real_normed_vector" shows "path_image(subpath u v g) = g ` (closed_segment u v)" by (auto simp add: closed_segment_real_eq path_image_def subpath_def) lemma path_image_subpath: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = (if u \ v then g ` {u..v} else g ` {v..u})" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_image_subpath_commute: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = path_image(subpath v u g)" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_subpath [simp]: fixes g :: "real \ 'a::real_normed_vector" assumes "path g" "u \ {0..1}" "v \ {0..1}" shows "path(subpath u v g)" proof - have "continuous_on {0..1} (g \ (\x. ((v-u) * x+ u)))" apply (rule continuous_intros | simp)+ apply (simp add: image_affinity_atLeastAtMost [where c=u]) using assms apply (auto simp: path_def continuous_on_subset) done then show ?thesis by (simp add: path_def subpath_def) qed lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)" by (simp add: pathstart_def subpath_def) lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)" by (simp add: pathfinish_def subpath_def) lemma subpath_trivial [simp]: "subpath 0 1 g = g" by (simp add: subpath_def) lemma subpath_reversepath: "subpath 1 0 g = reversepath g" by (simp add: reversepath_def subpath_def) lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g" by (simp add: reversepath_def subpath_def algebra_simps) lemma subpath_translation: "subpath u v ((\x. a + x) \ g) = (\x. a + x) \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma subpath_image: "subpath u v (f \ g) = f \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma affine_ineq: fixes x :: "'a::linordered_idom" assumes "x \ 1" "v \ u" shows "v + x * u \ u + x * v" proof - have "(1-x)*(u-v) \ 0" using assms by auto then show ?thesis by (simp add: algebra_simps) qed lemma sum_le_prod1: fixes a::real shows "\a \ 1; b \ 1\ \ a + b \ 1 + a * b" by (metis add.commute affine_ineq mult.right_neutral) lemma simple_path_subpath_eq: "simple_path(subpath u v g) \ path(subpath u v g) \ u\v \ (\x y. x \ closed_segment u v \ y \ closed_segment u v \ g x = g y \ x = y \ x = u \ y = v \ x = v \ y = u)" (is "?lhs = ?rhs") proof (rule iffI) assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" by (auto simp: simple_path_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y \ x = u \ y = v \ x = v \ y = u" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost) (simp_all add: field_split_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y \ x = u \ y = v \ x = v \ y = u" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y \ x = u \ y = v \ x = v \ y = u" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost) have [simp]: "\x. u + x * v = v + x * u \ u=v \ x=1" by algebra show ?lhs using psp ne unfolding simple_path_def subpath_def by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma arc_subpath_eq: "arc(subpath u v g) \ path(subpath u v g) \ u\v \ inj_on g (closed_segment u v)" (is "?lhs = ?rhs") proof (rule iffI) assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y)" by (auto simp: arc_def inj_on_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (cases "v = u") (simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost, simp add: field_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs unfolding inj_on_def by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost) show ?lhs using psp ne unfolding arc_def subpath_def inj_on_def by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma simple_path_subpath: assumes "simple_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "simple_path(subpath u v g)" using assms apply (simp add: simple_path_subpath_eq simple_path_imp_path) apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce) done lemma arc_simple_path_subpath: "\simple_path g; u \ {0..1}; v \ {0..1}; g u \ g v\ \ arc(subpath u v g)" by (force intro: simple_path_subpath simple_path_imp_arc) lemma arc_subpath_arc: "\arc g; u \ {0..1}; v \ {0..1}; u \ v\ \ arc(subpath u v g)" by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD) lemma arc_simple_path_subpath_interior: "\simple_path g; u \ {0..1}; v \ {0..1}; u \ v; \u-v\ < 1\ \ arc(subpath u v g)" apply (rule arc_simple_path_subpath) apply (force simp: simple_path_def)+ done lemma path_image_subpath_subset: "\u \ {0..1}; v \ {0..1}\ \ path_image(subpath u v g) \ path_image g" apply (simp add: closed_segment_real_eq image_affinity_atLeastAtMost path_image_subpath) apply (auto simp: path_image_def) done lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p" by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps) subsection\<^marker>\tag unimportant\\There is a subpath to the frontier\ lemma subpath_to_frontier_explicit: fixes S :: "'a::metric_space set" assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "\x. 0 \ x \ x < u \ g x \ interior S" "(g u \ interior S)" "(u = 0 \ g u \ closure S)" proof - have gcon: "continuous_on {0..1} g" using g by (simp add: path_def) then have com: "compact ({0..1} \ {u. g u \ closure (- S)})" apply (simp add: Int_commute [of "{0..1}"] compact_eq_bounded_closed closed_vimage_Int [unfolded vimage_def]) using compact_eq_bounded_closed apply fastforce done have "1 \ {u. g u \ closure (- S)}" using assms by (simp add: pathfinish_def closure_def) then have dis: "{0..1} \ {u. g u \ closure (- S)} \ {}" using atLeastAtMost_iff zero_le_one by blast then obtain u where "0 \ u" "u \ 1" and gu: "g u \ closure (- S)" and umin: "\t. \0 \ t; t \ 1; g t \ closure (- S)\ \ u \ t" using compact_attains_inf [OF com dis] by fastforce then have umin': "\t. \0 \ t; t \ 1; t < u\ \ g t \ S" using closure_def by fastforce { assume "u \ 0" then have "u > 0" using \0 \ u\ by auto { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {0..1}; dist x' u \ d\ \ dist (g x') (g u) < e" using continuous_onE [OF gcon _ \e > 0\] \0 \ _\ \_ \ 1\ atLeastAtMost_iff by auto have *: "dist (max 0 (u - d / 2)) u \ d" using \0 \ u\ \u \ 1\ \d > 0\ by (simp add: dist_real_def) have "\y\S. dist y (g u) < e" using \0 < u\ \u \ 1\ \d > 0\ by (force intro: d [OF _ *] umin') } then have "g u \ closure S" by (simp add: frontier_def closure_approachable) } then show ?thesis apply (rule_tac u=u in that) apply (auto simp: \0 \ u\ \u \ 1\ gu interior_closure umin) using \_ \ 1\ interior_closure umin apply fastforce done qed lemma subpath_to_frontier_strong: assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ interior S" "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" proof - obtain u where "0 \ u" "u \ 1" and gxin: "\x. 0 \ x \ x < u \ g x \ interior S" and gunot: "(g u \ interior S)" and u0: "(u = 0 \ g u \ closure S)" using subpath_to_frontier_explicit [OF assms] by blast show ?thesis apply (rule that [OF \0 \ u\ \u \ 1\]) apply (simp add: gunot) using \0 \ u\ u0 by (force simp: subpath_def gxin) qed lemma subpath_to_frontier: assumes g: "path g" and g0: "pathstart g \ closure S" and g1: "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" proof - obtain u where "0 \ u" "u \ 1" and notin: "g u \ interior S" and disj: "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" using subpath_to_frontier_strong [OF g g1] by blast show ?thesis apply (rule that [OF \0 \ u\ \u \ 1\]) apply (metis DiffI disj frontier_def g0 notin pathstart_def) using \0 \ u\ g0 disj apply (simp add: path_image_subpath_gen) apply (auto simp: closed_segment_eq_real_ivl pathstart_def pathfinish_def subpath_def) apply (rename_tac y) apply (drule_tac x="y/u" in spec) apply (auto split: if_split_asm) done qed lemma exists_path_subpath_to_frontier: fixes S :: "'a::real_normed_vector set" assumes "path g" "pathstart g \ closure S" "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" proof - obtain u where u: "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" using subpath_to_frontier [OF assms] by blast show ?thesis apply (rule that [of "subpath 0 u g"]) using assms u apply (simp_all add: path_image_subpath) apply (simp add: pathstart_def) apply (force simp: closed_segment_eq_real_ivl path_image_def) done qed lemma exists_path_subpath_to_frontier_closed: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and g: "path g" and g0: "pathstart g \ S" and g1: "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g \ S" "pathfinish h \ frontier S" proof - obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto show ?thesis apply (rule that [OF \path h\]) using assms h apply auto apply (metis Diff_single_insert frontier_subset_eq insert_iff interior_subset subset_iff) done qed subsection \Shift Path to Start at Some Given Point\ definition\<^marker>\tag important\ shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" lemma pathstart_shiftpath: "a \ 1 \ pathstart (shiftpath a g) = g a" unfolding pathstart_def shiftpath_def by auto lemma pathfinish_shiftpath: assumes "0 \ a" and "pathfinish g = pathstart g" shows "pathfinish (shiftpath a g) = g a" using assms unfolding pathstart_def pathfinish_def shiftpath_def by auto lemma endpoints_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0 .. 1}" shows "pathfinish (shiftpath a g) = g a" and "pathstart (shiftpath a g) = g a" using assms by (auto intro!: pathfinish_shiftpath pathstart_shiftpath) lemma closed_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)" using endpoints_shiftpath[OF assms] by auto lemma path_shiftpath: assumes "path g" and "pathfinish g = pathstart g" and "a \ {0..1}" shows "path (shiftpath a g)" proof - have *: "{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by auto have **: "\x. x + a = 1 \ g (x + a - 1) = g (x + a)" using assms(2)[unfolded pathfinish_def pathstart_def] by auto show ?thesis unfolding path_def shiftpath_def * proof (rule continuous_on_closed_Un) have contg: "continuous_on {0..1} g" using \path g\ path_def by blast show "continuous_on {0..1-a} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {0..1-a} (g \ (+) a)" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed auto show "continuous_on {1-a..1} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {1-a..1} (g \ (+) (a - 1))" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed (auto simp: "**" add.commute add_diff_eq) qed auto qed lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" and "x \ {0..1}" shows "shiftpath (1 - a) (shiftpath a g) x = g x" using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto lemma path_image_shiftpath: assumes a: "a \ {0..1}" and "pathfinish g = pathstart g" shows "path_image (shiftpath a g) = path_image g" proof - { fix x assume g: "g 1 = g 0" "x \ {0..1::real}" and gne: "\y. y\{0..1} \ {x. \ a + x \ 1} \ g x \ g (a + y - 1)" then have "\y\{0..1} \ {x. a + x \ 1}. g x = g (a + y)" proof (cases "a \ x") case False then show ?thesis apply (rule_tac x="1 + x - a" in bexI) using g gne[of "1 + x - a"] a apply (force simp: field_simps)+ done next case True then show ?thesis using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps) qed } then show ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def by (auto simp: image_iff) qed lemma simple_path_shiftpath: assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" shows "simple_path (shiftpath a g)" unfolding simple_path_def proof (intro conjI impI ballI) show "path (shiftpath a g)" by (simp add: assms path_shiftpath simple_path_imp_path) have *: "\x y. \g x = g y; x \ {0..1}; y \ {0..1}\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" if "x \ {0..1}" "y \ {0..1}" "shiftpath a g x = shiftpath a g y" for x y using that a unfolding shiftpath_def by (force split: if_split_asm dest!: *) qed subsection \Straight-Line Paths\ definition\<^marker>\tag important\ linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" where "linepath a b = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b)" lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a" unfolding pathstart_def linepath_def by auto lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b" unfolding pathfinish_def linepath_def by auto lemma linepath_inner: "linepath a b x \ v = linepath (a \ v) (b \ v) x" by (simp add: linepath_def algebra_simps) lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" by (simp add: linepath_def) lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" by (simp add: linepath_def) lemma linepath_0': "linepath a b 0 = a" by (simp add: linepath_def) lemma linepath_1': "linepath a b 1 = b" by (simp add: linepath_def) lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" unfolding linepath_def by (intro continuous_intros) lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" using continuous_linepath_at by (auto intro!: continuous_at_imp_continuous_on) lemma path_linepath[iff]: "path (linepath a b)" unfolding path_def by (rule continuous_on_linepath) lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b" unfolding path_image_def segment linepath_def by auto lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a" unfolding reversepath_def linepath_def by auto lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" by (simp add: linepath_def) lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma arc_linepath: assumes "a \ b" shows [simp]: "arc (linepath a b)" proof - { fix x y :: "real" assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b" then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps) with assms have "x = y" by simp } then show ?thesis unfolding arc_def inj_on_def by (fastforce simp: algebra_simps linepath_def) qed lemma simple_path_linepath[intro]: "a \ b \ simple_path (linepath a b)" by (simp add: arc_imp_simple_path) lemma linepath_trivial [simp]: "linepath a a x = a" by (simp add: linepath_def real_vector.scale_left_diff_distrib) lemma linepath_refl: "linepath a a = (\x. a)" by auto lemma subpath_refl: "subpath a a g = linepath (g a) (g a)" by (simp add: subpath_def linepath_def algebra_simps) lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)" by (simp add: scaleR_conv_of_real linepath_def) lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x" by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def) lemma inj_on_linepath: assumes "a \ b" shows "inj_on (linepath a b) {0..1}" proof (clarsimp simp: inj_on_def linepath_def) fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)" by (auto simp: algebra_simps) then show "x=y" using assms by auto qed lemma linepath_le_1: fixes a::"'a::linordered_idom" shows "\a \ 1; b \ 1; 0 \ u; u \ 1\ \ (1 - u) * a + u * b \ 1" using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto subsection\<^marker>\tag unimportant\\Segments via convex hulls\ lemma segments_subset_convex_hull: "closed_segment a b \ (convex hull {a,b,c})" "closed_segment a c \ (convex hull {a,b,c})" "closed_segment b c \ (convex hull {a,b,c})" "closed_segment b a \ (convex hull {a,b,c})" "closed_segment c a \ (convex hull {a,b,c})" "closed_segment c b \ (convex hull {a,b,c})" by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono]) lemma midpoints_in_convex_hull: assumes "x \ convex hull s" "y \ convex hull s" shows "midpoint x y \ convex hull s" proof - have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \ convex hull s" by (rule convexD_alt) (use assms in auto) then show ?thesis by (simp add: midpoint_def algebra_simps) qed lemma not_in_interior_convex_hull_3: fixes a :: "complex" shows "a \ interior(convex hull {a,b,c})" "b \ interior(convex hull {a,b,c})" "c \ interior(convex hull {a,b,c})" by (auto simp: card_insert_le_m1 not_in_interior_convex_hull) lemma midpoint_in_closed_segment [simp]: "midpoint a b \ closed_segment a b" using midpoints_in_convex_hull segment_convex_hull by blast lemma midpoint_in_open_segment [simp]: "midpoint a b \ open_segment a b \ a \ b" by (simp add: open_segment_def) lemma continuous_IVT_local_extremum: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and "a \ b" "f a = f b" obtains z where "z \ open_segment a b" "(\w \ closed_segment a b. (f w) \ (f z)) \ (\w \ closed_segment a b. (f z) \ (f w))" proof - obtain c where "c \ closed_segment a b" and c: "\y. y \ closed_segment a b \ f y \ f c" using continuous_attains_sup [of "closed_segment a b" f] contf by auto obtain d where "d \ closed_segment a b" and d: "\y. y \ closed_segment a b \ f d \ f y" using continuous_attains_inf [of "closed_segment a b" f] contf by auto show ?thesis proof (cases "c \ open_segment a b \ d \ open_segment a b") case True then show ?thesis using c d that by blast next case False then have "(c = a \ c = b) \ (d = a \ d = b)" by (simp add: \c \ closed_segment a b\ \d \ closed_segment a b\ open_segment_def) with \a \ b\ \f a = f b\ c d show ?thesis by (rule_tac z = "midpoint a b" in that) (fastforce+) qed qed text\An injective map into R is also an open map w.r.T. the universe, and conversely. \ proposition injective_eq_1d_open_map_UNIV: fixes f :: "real \ real" assumes contf: "continuous_on S f" and S: "is_interval S" shows "inj_on f S \ (\T. open T \ T \ S \ open(f ` T))" (is "?lhs = ?rhs") proof safe fix T assume injf: ?lhs and "open T" and "T \ S" have "\U. open U \ f x \ U \ U \ f ` T" if "x \ T" for x proof - obtain \ where "\ > 0" and \: "cball x \ \ T" using \open T\ \x \ T\ open_contains_cball_eq by blast show ?thesis proof (intro exI conjI) have "closed_segment (x-\) (x+\) = {x-\..x+\}" using \0 < \\ by (auto simp: closed_segment_eq_real_ivl) also have "\ \ S" using \ \T \ S\ by (auto simp: dist_norm subset_eq) finally have "f ` (open_segment (x-\) (x+\)) = open_segment (f (x-\)) (f (x+\))" using continuous_injective_image_open_segment_1 by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf]) then show "open (f ` {x-\<..})" using \0 < \\ by (simp add: open_segment_eq_real_ivl) show "f x \ f ` {x - \<..}" by (auto simp: \\ > 0\) show "f ` {x - \<..} \ f ` T" using \ by (auto simp: dist_norm subset_iff) qed qed with open_subopen show "open (f ` T)" by blast next assume R: ?rhs have False if xy: "x \ S" "y \ S" and "f x = f y" "x \ y" for x y proof - have "open (f ` open_segment x y)" using R by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy) moreover have "continuous_on (closed_segment x y) f" by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that) then obtain \ where "\ \ open_segment x y" and \: "(\w \ closed_segment x y. (f w) \ (f \)) \ (\w \ closed_segment x y. (f \) \ (f w))" using continuous_IVT_local_extremum [of x y f] \f x = f y\ \x \ y\ by blast ultimately obtain e where "e>0" and e: "\u. dist u (f \) < e \ u \ f ` open_segment x y" using open_dist by (metis image_eqI) have fin: "f \ + (e/2) \ f ` open_segment x y" "f \ - (e/2) \ f ` open_segment x y" using e [of "f \ + (e/2)"] e [of "f \ - (e/2)"] \e > 0\ by (auto simp: dist_norm) show ?thesis using \ \0 < e\ fin open_closed_segment by fastforce qed then show ?lhs by (force simp: inj_on_def) qed subsection\<^marker>\tag unimportant\ \Bounding a point away from a path\ lemma not_on_path_ball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and z: "z \ path_image g" shows "\e > 0. ball z e \ path_image g = {}" proof - have "closed (path_image g)" by (simp add: \path g\ closed_path_image) then obtain a where "a \ path_image g" "\y \ path_image g. dist z a \ dist z y" by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z]) then show ?thesis by (rule_tac x="dist z a" in exI) (use dist_commute z in auto) qed lemma not_on_path_cball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and "z \ path_image g" shows "\e>0. cball z e \ (path_image g) = {}" proof - obtain e where "ball z e \ path_image g = {}" "e > 0" using not_on_path_ball[OF assms] by auto moreover have "cball z (e/2) \ ball z e" using \e > 0\ by auto ultimately show ?thesis by (rule_tac x="e/2" in exI) auto qed subsection \Path component\ text \Original formalization by Tom Hales\ definition\<^marker>\tag important\ "path_component s x y \ (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" abbreviation\<^marker>\tag important\ "path_component_set s x \ Collect (path_component s x)" lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def lemma path_component_mem: assumes "path_component s x y" shows "x \ s" and "y \ s" using assms unfolding path_defs by auto lemma path_component_refl: assumes "x \ s" shows "path_component s x x" unfolding path_defs apply (rule_tac x="\u. x" in exI) using assms apply (auto intro!: continuous_intros) done lemma path_component_refl_eq: "path_component s x x \ x \ s" by (auto intro!: path_component_mem path_component_refl) lemma path_component_sym: "path_component s x y \ path_component s y x" unfolding path_component_def apply (erule exE) apply (rule_tac x="reversepath g" in exI, auto) done lemma path_component_trans: assumes "path_component s x y" and "path_component s y z" shows "path_component s x z" using assms unfolding path_component_def apply (elim exE) apply (rule_tac x="g +++ ga" in exI) apply (auto simp: path_image_join) done lemma path_component_of_subset: "s \ t \ path_component s x y \ path_component t x y" unfolding path_component_def by auto lemma path_component_linepath: fixes s :: "'a::real_normed_vector set" shows "closed_segment a b \ s \ path_component s a b" unfolding path_component_def by (rule_tac x="linepath a b" in exI, auto) subsubsection\<^marker>\tag unimportant\ \Path components as sets\ lemma path_component_set: "path_component_set s x = {y. (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)}" by (auto simp: path_component_def) lemma path_component_subset: "path_component_set s x \ s" by (auto simp: path_component_mem(2)) lemma path_component_eq_empty: "path_component_set s x = {} \ x \ s" using path_component_mem path_component_refl_eq by fastforce lemma path_component_mono: "s \ t \ (path_component_set s x) \ (path_component_set t x)" by (simp add: Collect_mono path_component_of_subset) lemma path_component_eq: "y \ path_component_set s x \ path_component_set s y = path_component_set s x" by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans) subsection \Path connectedness of a space\ definition\<^marker>\tag important\ "path_connected s \ (\x\s. \y\s. \g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" lemma path_connectedin_iff_path_connected_real [simp]: "path_connectedin euclideanreal S \ path_connected S" by (simp add: path_connectedin path_connected_def path_defs) lemma path_connected_component: "path_connected s \ (\x\s. \y\s. path_component s x y)" unfolding path_connected_def path_component_def by auto lemma path_connected_component_set: "path_connected s \ (\x\s. path_component_set s x = s)" unfolding path_connected_component path_component_subset using path_component_mem by blast lemma path_component_maximal: "\x \ t; path_connected t; t \ s\ \ t \ (path_component_set s x)" by (metis path_component_mono path_connected_component_set) lemma convex_imp_path_connected: fixes s :: "'a::real_normed_vector set" assumes "convex s" shows "path_connected s" unfolding path_connected_def using assms convex_contains_segment by fastforce lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)" by (simp add: convex_imp_path_connected) lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)" using path_connected_component_set by auto lemma path_connected_imp_connected: assumes "path_connected S" shows "connected S" proof (rule connectedI) fix e1 e2 assume as: "open e1" "open e2" "S \ e1 \ e2" "e1 \ e2 \ S = {}" "e1 \ S \ {}" "e2 \ S \ {}" then obtain x1 x2 where obt:"x1 \ e1 \ S" "x2 \ e2 \ S" by auto then obtain g where g: "path g" "path_image g \ S" "pathstart g = x1" "pathfinish g = x2" using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto have *: "connected {0..1::real}" by (auto intro!: convex_connected) have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl) ultimately show False using *[unfolded connected_local not_ex, rule_format, of "{0..1} \ g -` e1" "{0..1} \ g -` e2"] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)] by auto qed lemma open_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (path_component_set S x)" unfolding open_contains_ball proof fix y assume as: "y \ path_component_set S x" then have "y \ S" by (simp add: path_component_mem(2)) then obtain e where e: "e > 0" "ball y e \ S" using assms[unfolded open_contains_ball] by auto have "\u. dist y u < e \ path_component S x u" by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component) then show "\e > 0. ball y e \ path_component_set S x" using \e>0\ by auto qed lemma open_non_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (S - path_component_set S x)" unfolding open_contains_ball proof fix y assume y: "y \ S - path_component_set S x" then obtain e where e: "e > 0" "ball y e \ S" using assms openE by auto show "\e>0. ball y e \ S - path_component_set S x" proof (intro exI conjI subsetI DiffI notI) show "\x. x \ ball y e \ x \ S" using e by blast show False if "z \ ball y e" "z \ path_component_set S x" for z proof - have "y \ path_component_set S z" by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1)) then have "y \ path_component_set S x" using path_component_eq that(2) by blast then show False using y by blast qed qed (use e in auto) qed lemma connected_open_path_connected: fixes S :: "'a::real_normed_vector set" assumes "open S" and "connected S" shows "path_connected S" unfolding path_connected_component_set proof (rule, rule, rule path_component_subset, rule) fix x y assume "x \ S" and "y \ S" show "y \ path_component_set S x" proof (rule ccontr) assume "\ ?thesis" moreover have "path_component_set S x \ S \ {}" using \x \ S\ path_component_eq_empty path_component_subset[of S x] by auto ultimately show False using \y \ S\ open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] using assms(2)[unfolded connected_def not_ex, rule_format, of "path_component_set S x" "S - path_component_set S x"] by auto qed qed lemma path_connected_continuous_image: assumes "continuous_on S f" and "path_connected S" shows "path_connected (f ` S)" unfolding path_connected_def proof (rule, rule) fix x' y' assume "x' \ f ` S" "y' \ f ` S" then obtain x y where x: "x \ S" and y: "y \ S" and x': "x' = f x" and y': "y' = f y" by auto from x y obtain g where "path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" using assms(2)[unfolded path_connected_def] by fast then show "\g. path g \ path_image g \ f ` S \ pathstart g = x' \ pathfinish g = y'" unfolding x' y' apply (rule_tac x="f \ g" in exI) unfolding path_defs apply (intro conjI continuous_on_compose continuous_on_subset[OF assms(1)]) apply auto done qed lemma path_connected_translationI: fixes a :: "'a :: topological_group_add" assumes "path_connected S" shows "path_connected ((\x. a + x) ` S)" by (intro path_connected_continuous_image assms continuous_intros) lemma path_connected_translation: fixes a :: "'a :: topological_group_add" shows "path_connected ((\x. a + x) ` S) = path_connected S" proof - have "\x y. (+) (x::'a) ` (+) (0 - x) ` y = y" by (simp add: image_image) then show ?thesis by (metis (no_types) path_connected_translationI) qed lemma path_connected_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (closed_segment a b)" by (simp add: convex_imp_path_connected) lemma path_connected_open_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (open_segment a b)" by (simp add: convex_imp_path_connected) lemma homeomorphic_path_connectedness: "S homeomorphic T \ path_connected S \ path_connected T" unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) lemma path_connected_empty [simp]: "path_connected {}" unfolding path_connected_def by auto lemma path_connected_singleton [simp]: "path_connected {a}" unfolding path_connected_def pathstart_def pathfinish_def path_image_def apply clarify apply (rule_tac x="\x. a" in exI) apply (simp add: image_constant_conv) apply (simp add: path_def) done lemma path_connected_Un: assumes "path_connected S" and "path_connected T" and "S \ T \ {}" shows "path_connected (S \ T)" unfolding path_connected_component proof (intro ballI) fix x y assume x: "x \ S \ T" and y: "y \ S \ T" from assms obtain z where z: "z \ S" "z \ T" by auto show "path_component (S \ T) x y" using x y proof safe assume "x \ S" "y \ S" then show "path_component (S \ T) x y" by (meson Un_upper1 \path_connected S\ path_component_of_subset path_connected_component) next assume "x \ S" "y \ T" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ S" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ T" then show "path_component (S \ T) x y" by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute) qed qed lemma path_connected_UNION: assumes "\i. i \ A \ path_connected (S i)" and "\i. i \ A \ z \ S i" shows "path_connected (\i\A. S i)" unfolding path_connected_component proof clarify fix x i y j assume *: "i \ A" "x \ S i" "j \ A" "y \ S j" then have "path_component (S i) x z" and "path_component (S j) z y" using assms by (simp_all add: path_connected_component) then have "path_component (\i\A. S i) x z" and "path_component (\i\A. S i) z y" using *(1,3) by (auto elim!: path_component_of_subset [rotated]) then show "path_component (\i\A. S i) x y" by (rule path_component_trans) qed lemma path_component_path_image_pathstart: assumes p: "path p" and x: "x \ path_image p" shows "path_component (path_image p) (pathstart p) x" proof - obtain y where x: "x = p y" and y: "0 \ y" "y \ 1" using x by (auto simp: path_image_def) show ?thesis unfolding path_component_def proof (intro exI conjI) have "continuous_on {0..1} (p \ ((*) y))" apply (rule continuous_intros)+ using p [unfolded path_def] y apply (auto simp: mult_le_one intro: continuous_on_subset [of _ p]) done then show "path (\u. p (y * u))" by (simp add: path_def) show "path_image (\u. p (y * u)) \ path_image p" using y mult_le_one by (fastforce simp: path_image_def image_iff) qed (auto simp: pathstart_def pathfinish_def x) qed lemma path_connected_path_image: "path p \ path_connected(path_image p)" unfolding path_connected_component by (meson path_component_path_image_pathstart path_component_sym path_component_trans) lemma path_connected_path_component [simp]: "path_connected (path_component_set s x)" proof - { fix y z assume pa: "path_component s x y" "path_component s x z" then have pae: "path_component_set s x = path_component_set s y" using path_component_eq by auto have yz: "path_component s y z" using pa path_component_sym path_component_trans by blast then have "\g. path g \ path_image g \ path_component_set s x \ pathstart g = y \ pathfinish g = z" apply (simp add: path_component_def, clarify) apply (rule_tac x=g in exI) by (simp add: pae path_component_maximal path_connected_path_image pathstart_in_path_image) } then show ?thesis by (simp add: path_connected_def) qed lemma path_component: "path_component S x y \ (\t. path_connected t \ t \ S \ x \ t \ y \ t)" apply (intro iffI) apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image) using path_component_of_subset path_connected_component by blast lemma path_component_path_component [simp]: "path_component_set (path_component_set S x) x = path_component_set S x" proof (cases "x \ S") case True show ?thesis apply (rule subset_antisym) apply (simp add: path_component_subset) by (simp add: True path_component_maximal path_component_refl) next case False then show ?thesis by (metis False empty_iff path_component_eq_empty) qed lemma path_component_subset_connected_component: "(path_component_set S x) \ (connected_component_set S x)" proof (cases "x \ S") case True show ?thesis apply (rule connected_component_maximal) apply (auto simp: True path_component_subset path_component_refl path_connected_imp_connected) done next case False then show ?thesis using path_component_eq_empty by auto qed subsection\<^marker>\tag unimportant\\Lemmas about path-connectedness\ lemma path_connected_linear_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "path_connected S" "bounded_linear f" shows "path_connected(f ` S)" by (auto simp: linear_continuous_on assms path_connected_continuous_image) lemma is_interval_path_connected: "is_interval S \ path_connected S" by (simp add: convex_imp_path_connected is_interval_convex) lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Iio[simp]: "path_connected {.. (\x \ topspace X. \y \ topspace X. \S. path_connectedin X S \ x \ S \ y \ S)" unfolding path_connected_space_def Ball_def apply (intro all_cong1 imp_cong refl, safe) using path_connectedin_path_image apply fastforce by (meson path_connectedin) lemma connectedin_path_image: "pathin X g \ connectedin X (g ` ({0..1}))" by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image) lemma compactin_path_image: "pathin X g \ compactin X (g ` ({0..1}))" unfolding pathin_def by (rule image_compactin [of "top_of_set {0..1}"]) auto lemma linear_homeomorphism_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" obtains g where "homeomorphism (f ` S) S g f" using linear_injective_left_inverse [OF assms] apply clarify apply (rule_tac g=g in that) using assms apply (auto simp: homeomorphism_def eq_id_iff [symmetric] image_comp comp_def linear_conv_bounded_linear linear_continuous_on) done lemma linear_homeomorphic_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "S homeomorphic f ` S" by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms]) lemma path_connected_Times: assumes "path_connected s" "path_connected t" shows "path_connected (s \ t)" proof (simp add: path_connected_def Sigma_def, clarify) fix x1 y1 x2 y2 assume "x1 \ s" "y1 \ t" "x2 \ s" "y2 \ t" obtain g where "path g" and g: "path_image g \ s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2" using \x1 \ s\ \x2 \ s\ assms by (force simp: path_connected_def) obtain h where "path h" and h: "path_image h \ t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2" using \y1 \ t\ \y2 \ t\ assms by (force simp: path_connected_def) have "path (\z. (x1, h z))" using \path h\ apply (simp add: path_def) apply (rule continuous_on_compose2 [where f = h]) apply (rule continuous_intros | force)+ done moreover have "path (\z. (g z, y2))" using \path g\ apply (simp add: path_def) apply (rule continuous_on_compose2 [where f = g]) apply (rule continuous_intros | force)+ done ultimately have 1: "path ((\z. (x1, h z)) +++ (\z. (g z, y2)))" by (metis hf gs path_join_imp pathstart_def pathfinish_def) have "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ path_image (\z. (x1, h z)) \ path_image (\z. (g z, y2))" by (rule Path_Connected.path_image_join_subset) also have "\ \ (\x\s. \x1\t. {(x, x1)})" using g h \x1 \ s\ \y2 \ t\ by (force simp: path_image_def) finally have 2: "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ (\x\s. \x1\t. {(x, x1)})" . show "\g. path g \ path_image g \ (\x\s. \x1\t. {(x, x1)}) \ pathstart g = (x1, y1) \ pathfinish g = (x2, y2)" apply (intro exI conjI) apply (rule 1) apply (rule 2) apply (metis hs pathstart_def pathstart_join) by (metis gf pathfinish_def pathfinish_join) qed lemma is_interval_path_connected_1: fixes s :: "real set" shows "is_interval s \ path_connected s" using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast subsection\<^marker>\tag unimportant\\Path components\ lemma Union_path_component [simp]: "Union {path_component_set S x |x. x \ S} = S" apply (rule subset_antisym) using path_component_subset apply force using path_component_refl by auto lemma path_component_disjoint: "disjnt (path_component_set S a) (path_component_set S b) \ (a \ path_component_set S b)" apply (auto simp: disjnt_def) using path_component_eq apply fastforce using path_component_sym path_component_trans by blast lemma path_component_eq_eq: "path_component S x = path_component S y \ (x \ S) \ (y \ S) \ x \ S \ y \ S \ path_component S x y" apply (rule iffI, metis (no_types) path_component_mem(1) path_component_refl) apply (erule disjE, metis Collect_empty_eq_bot path_component_eq_empty) apply (rule ext) apply (metis path_component_trans path_component_sym) done lemma path_component_unique: assumes "x \ c" "c \ S" "path_connected c" "\c'. \x \ c'; c' \ S; path_connected c'\ \ c' \ c" shows "path_component_set S x = c" apply (rule subset_antisym) using assms apply (metis mem_Collect_eq subsetCE path_component_eq_eq path_component_subset path_connected_path_component) by (simp add: assms path_component_maximal) lemma path_component_intermediate_subset: "path_component_set u a \ t \ t \ u \ path_component_set t a = path_component_set u a" by (metis (no_types) path_component_mono path_component_path_component subset_antisym) lemma complement_path_component_Union: fixes x :: "'a :: topological_space" shows "S - path_component_set S x = \({path_component_set S y| y. y \ S} - {path_component_set S x})" proof - have *: "(\x. x \ S - {a} \ disjnt a x) \ \S - a = \(S - {a})" for a::"'a set" and S by (auto simp: disjnt_def) have "\y. y \ {path_component_set S x |x. x \ S} - {path_component_set S x} \ disjnt (path_component_set S x) y" using path_component_disjoint path_component_eq by fastforce then have "\{path_component_set S x |x. x \ S} - path_component_set S x = \({path_component_set S y |y. y \ S} - {path_component_set S x})" by (meson *) then show ?thesis by simp qed subsection\Path components\ definition path_component_of where "path_component_of X x y \ \g. pathin X g \ g 0 = x \ g 1 = y" abbreviation path_component_of_set where "path_component_of_set X x \ Collect (path_component_of X x)" definition path_components_of :: "'a topology \ 'a set set" where "path_components_of X \ path_component_of_set X ` topspace X" lemma pathin_canon_iff: "pathin (top_of_set T) g \ path g \ g ` {0..1} \ T" by (simp add: path_def pathin_def) lemma path_component_of_canon_iff [simp]: "path_component_of (top_of_set T) a b \ path_component T a b" by (simp add: path_component_of_def pathin_canon_iff path_defs) lemma path_component_in_topspace: "path_component_of X x y \ x \ topspace X \ y \ topspace X" by (auto simp: path_component_of_def pathin_def continuous_map_def) lemma path_component_of_refl: "path_component_of X x x \ x \ topspace X" apply (auto simp: path_component_in_topspace) apply (force simp: path_component_of_def pathin_const) done lemma path_component_of_sym: assumes "path_component_of X x y" shows "path_component_of X y x" using assms apply (clarsimp simp: path_component_of_def pathin_def) apply (rule_tac x="g \ (\t. 1 - t)" in exI) apply (auto intro!: continuous_map_compose) apply (force simp: continuous_map_in_subtopology continuous_on_op_minus) done lemma path_component_of_sym_iff: "path_component_of X x y \ path_component_of X y x" by (metis path_component_of_sym) lemma path_component_of_trans: assumes "path_component_of X x y" and "path_component_of X y z" shows "path_component_of X x z" unfolding path_component_of_def pathin_def proof - let ?T01 = "top_of_set {0..1::real}" obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1" and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1" using assms unfolding path_component_of_def pathin_def by blast let ?g = "\x. if x \ 1/2 then (g1 \ (\t. 2 * t)) x else (g2 \ (\t. 2 * t -1)) x" show "\g. continuous_map ?T01 X g \ g 0 = x \ g 1 = z" proof (intro exI conjI) show "continuous_map (subtopology euclideanreal {0..1}) X ?g" proof (intro continuous_map_cases_le continuous_map_compose, force, force) show "continuous_map (subtopology ?T01 {x \ topspace ?T01. x \ 1/2}) ?T01 ((*) 2)" by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology) have "continuous_map (subtopology (top_of_set {0..1}) {x. 0 \ x \ x \ 1 \ 1 \ x * 2}) euclideanreal (\t. 2 * t - 1)" by (intro continuous_intros) (force intro: continuous_map_from_subtopology) then show "continuous_map (subtopology ?T01 {x \ topspace ?T01. 1/2 \ x}) ?T01 (\t. 2 * t - 1)" by (force simp: continuous_map_in_subtopology) show "(g1 \ (*) 2) x = (g2 \ (\t. 2 * t - 1)) x" if "x \ topspace ?T01" "x = 1/2" for x using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology) qed (auto simp: g1 g2) qed (auto simp: g1 g2) qed lemma path_component_of_mono: "\path_component_of (subtopology X S) x y; S \ T\ \ path_component_of (subtopology X T) x y" unfolding path_component_of_def by (metis subsetD pathin_subtopology) lemma path_component_of: "path_component_of X x y \ (\T. path_connectedin X T \ x \ T \ y \ T)" apply (auto simp: path_component_of_def) using path_connectedin_path_image apply fastforce apply (metis path_connectedin) done lemma path_component_of_set: "path_component_of X x y \ (\g. pathin X g \ g 0 = x \ g 1 = y)" by (auto simp: path_component_of_def) lemma path_component_of_subset_topspace: "Collect(path_component_of X x) \ topspace X" using path_component_in_topspace by fastforce lemma path_component_of_eq_empty: "Collect(path_component_of X x) = {} \ (x \ topspace X)" using path_component_in_topspace path_component_of_refl by fastforce lemma path_connected_space_iff_path_component: "path_connected_space X \ (\x \ topspace X. \y \ topspace X. path_component_of X x y)" by (simp add: path_component_of path_connected_space_subconnected) lemma path_connected_space_imp_path_component_of: "\path_connected_space X; a \ topspace X; b \ topspace X\ \ path_component_of X a b" by (simp add: path_connected_space_iff_path_component) lemma path_connected_space_path_component_set: "path_connected_space X \ (\x \ topspace X. Collect(path_component_of X x) = topspace X)" using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce lemma path_component_of_maximal: "\path_connectedin X s; x \ s\ \ s \ Collect(path_component_of X x)" using path_component_of by fastforce lemma path_component_of_equiv: "path_component_of X x y \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: fun_eq_iff path_component_in_topspace) apply (meson path_component_of_sym path_component_of_trans) done qed (simp add: path_component_of_refl) lemma path_component_of_disjoint: "disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \ ~(path_component_of X x y)" by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_eq: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ path_component_of X x y" by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv) lemma path_connectedin_path_component_of: "path_connectedin X (Collect (path_component_of X x))" proof - have "\y. path_component_of X x y \ path_component_of (subtopology X (Collect (path_component_of X x))) x y" by (meson path_component_of path_component_of_maximal path_connectedin_subtopology) then show ?thesis apply (simp add: path_connectedin_def path_component_of_subset_topspace path_connected_space_iff_path_component) by (metis Int_absorb1 mem_Collect_eq path_component_of_equiv path_component_of_subset_topspace topspace_subtopology) qed lemma path_connectedin_euclidean [simp]: "path_connectedin euclidean S \ path_connected S" by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component) lemma path_connected_space_euclidean_subtopology [simp]: "path_connected_space(subtopology euclidean S) \ path_connected S" using path_connectedin_topspace by force lemma Union_path_components_of: "\(path_components_of X) = topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_components_of_maximal: "\C \ path_components_of X; path_connectedin X S; ~disjnt C S\ \ S \ C" apply (auto simp: path_components_of_def path_component_of_equiv) using path_component_of_maximal path_connectedin_def apply fastforce by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal) lemma pairwise_disjoint_path_components_of: "pairwise disjnt (path_components_of X)" by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv) lemma complement_path_components_of_Union: "C \ path_components_of X \ topspace X - C = \(path_components_of X - {C})" by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of) lemma nonempty_path_components_of: "C \ path_components_of X \ (C \ {})" apply (clarsimp simp: path_components_of_def path_component_of_eq_empty) by (meson path_component_of_refl) lemma path_components_of_subset: "C \ path_components_of X \ C \ topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_connectedin_path_components_of: "C \ path_components_of X \ path_connectedin X C" by (auto simp: path_components_of_def path_connectedin_path_component_of) lemma path_component_in_path_components_of: "Collect (path_component_of X a) \ path_components_of X \ a \ topspace X" apply (rule iffI) using nonempty_path_components_of path_component_of_eq_empty apply fastforce by (simp add: path_components_of_def) lemma path_connectedin_Union: assumes \: "\S. S \ \ \ path_connectedin X S" "\\ \ {}" shows "path_connectedin X (\\)" proof - obtain a where "\S. S \ \ \ a \ S" using assms by blast then have "\x. x \ topspace (subtopology X (\\)) \ path_component_of (subtopology X (\\)) a x" apply (simp) by (meson Union_upper \ path_component_of path_connectedin_subtopology) then show ?thesis using \ unfolding path_connectedin_def by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component) qed lemma path_connectedin_Un: "\path_connectedin X S; path_connectedin X T; S \ T \ {}\ \ path_connectedin X (S \ T)" by (blast intro: path_connectedin_Union [of "{S,T}", simplified]) lemma path_connected_space_iff_components_eq: "path_connected_space X \ (\C \ path_components_of X. \C' \ path_components_of X. C = C')" unfolding path_components_of_def proof (intro iffI ballI) assume "\C \ path_component_of_set X ` topspace X. \C' \ path_component_of_set X ` topspace X. C = C'" then show "path_connected_space X" using path_component_of_refl path_connected_space_iff_path_component by fastforce qed (auto simp: path_connected_space_path_component_set) lemma path_components_of_eq_empty: "path_components_of X = {} \ topspace X = {}" using Union_path_components_of nonempty_path_components_of by fastforce lemma path_components_of_empty_space: "topspace X = {} \ path_components_of X = {}" by (simp add: path_components_of_eq_empty) lemma path_components_of_subset_singleton: "path_components_of X \ {S} \ path_connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty) next case False have "(path_components_of X = {S}) \ (path_connected_space X \ topspace X = S)" proof (intro iffI conjI) assume L: "path_components_of X = {S}" then show "path_connected_space X" by (simp add: path_connected_space_iff_components_eq) show "topspace X = S" by (metis L ccpo_Sup_singleton [of S] Union_path_components_of) next assume R: "path_connected_space X \ topspace X = S" then show "path_components_of X = {S}" using ccpo_Sup_singleton [of S] by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set) qed with False show ?thesis by (simp add: path_components_of_eq_empty subset_singleton_iff) qed lemma path_connected_space_iff_components_subset_singleton: "path_connected_space X \ (\a. path_components_of X \ {a})" by (simp add: path_components_of_subset_singleton) lemma path_components_of_eq_singleton: "path_components_of X = {S} \ path_connected_space X \ topspace X \ {} \ S = topspace X" by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff) lemma path_components_of_path_connected_space: "path_connected_space X \ path_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: path_components_of_eq_empty path_components_of_eq_singleton) lemma path_component_subset_connected_component_of: "path_component_of_set X x \ connected_component_of_set X x" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of) next case False then show ?thesis using path_component_of_eq_empty by fastforce qed lemma exists_path_component_of_superset: assumes S: "path_connectedin X S" and ne: "topspace X \ {}" obtains C where "C \ path_components_of X" "S \ C" proof (cases "S = {}") case True then show ?thesis using ne path_components_of_eq_empty that by fastforce next case False then obtain a where "a \ S" by blast show ?thesis proof show "Collect (path_component_of X a) \ path_components_of X" by (meson \a \ S\ S subsetD path_component_in_path_components_of path_connectedin_subset_topspace) show "S \ Collect (path_component_of X a)" by (simp add: S \a \ S\ path_component_of_maximal) qed qed lemma path_component_of_eq_overlap: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ Collect (path_component_of X x) \ Collect (path_component_of X y) \ {}" by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty) lemma path_component_of_nonoverlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) = {} \ (x \ topspace X) \ (y \ topspace X) \ path_component_of X x \ path_component_of X y" by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap) lemma path_component_of_overlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) \ {} \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" by (meson path_component_of_nonoverlap) lemma path_components_of_disjoint: "\C \ path_components_of X; C' \ path_components_of X\ \ disjnt C C' \ C \ C'" by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv) lemma path_components_of_overlap: "\C \ path_components_of X; C' \ path_components_of X\ \ C \ C' \ {} \ C = C'" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_component_of_unique: "\x \ C; path_connectedin X C; \C'. \x \ C'; path_connectedin X C'\ \ C' \ C\ \ Collect (path_component_of X x) = C" by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of) lemma path_component_of_discrete_topology [simp]: "Collect (path_component_of (discrete_topology U) x) = (if x \ U then {x} else {})" proof - have "\C'. \x \ C'; path_connectedin (discrete_topology U) C'\ \ C' \ {x}" by (metis path_connectedin_discrete_topology subsetD singletonD) then have "x \ U \ Collect (path_component_of (discrete_topology U) x) = {x}" by (simp add: path_component_of_unique) then show ?thesis using path_component_in_topspace by fastforce qed lemma path_component_of_discrete_topology_iff [simp]: "path_component_of (discrete_topology U) x y \ x \ U \ y=x" by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD) lemma path_components_of_discrete_topology [simp]: "path_components_of (discrete_topology U) = (\x. {x}) ` U" by (auto simp: path_components_of_def image_def fun_eq_iff) lemma homeomorphic_map_path_component_of: assumes f: "homeomorphic_map X Y f" and x: "x \ topspace X" shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)" proof - obtain g where g: "homeomorphic_maps X Y f g" using f homeomorphic_map_maps by blast show ?thesis proof have "Collect (path_component_of Y (f x)) \ topspace Y" by (simp add: path_component_of_subset_topspace) moreover have "g ` Collect(path_component_of Y (f x)) \ Collect (path_component_of X (g (f x)))" using g x unfolding homeomorphic_maps_def by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of) ultimately show "Collect (path_component_of Y (f x)) \ f ` Collect (path_component_of X x)" using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff by metis show "f ` Collect (path_component_of X x) \ Collect (path_component_of Y (f x))" proof (rule path_component_of_maximal) show "path_connectedin Y (f ` Collect (path_component_of X x))" by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of) qed (simp add: path_component_of_refl x) qed qed lemma homeomorphic_map_path_components_of: assumes "homeomorphic_map X Y f" shows "path_components_of Y = (image f) ` (path_components_of X)" unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric] apply safe using assms homeomorphic_map_path_component_of apply fastforce by (metis assms homeomorphic_map_path_component_of imageI) subsection \Sphere is path-connected\ lemma path_connected_punctured_universe: assumes "2 \ DIM('a::euclidean_space)" shows "path_connected (- {a::'a})" proof - let ?A = "{x::'a. \i\Basis. x \ i < a \ i}" let ?B = "{x::'a. \i\Basis. a \ i < x \ i}" have A: "path_connected ?A" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i - 1)*\<^sub>R i) \ {x::'a. x \ i < a \ i}" by simp show "path_connected {x. x \ i < a \ i}" using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \ i"] by (simp add: inner_commute) qed have B: "path_connected ?B" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i + 1) *\<^sub>R i) \ {x::'a. a \ i < x \ i}" by simp show "path_connected {x. a \ i < x \ i}" using convex_imp_path_connected [OF convex_halfspace_gt, of "a \ i" i] by (simp add: inner_commute) qed obtain S :: "'a set" where "S \ Basis" and "card S = Suc (Suc 0)" using ex_card[OF assms] by auto then obtain b0 b1 :: 'a where "b0 \ Basis" and "b1 \ Basis" and "b0 \ b1" unfolding card_Suc_eq by auto then have "a + b0 - b1 \ ?A \ ?B" by (auto simp: inner_simps inner_Basis) then have "?A \ ?B \ {}" by fast with A B have "path_connected (?A \ ?B)" by (rule path_connected_Un) also have "?A \ ?B = {x. \i\Basis. x \ i \ a \ i}" unfolding neq_iff bex_disj_distrib Collect_disj_eq .. also have "\ = {x. x \ a}" unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def) also have "\ = - {a}" by auto finally show ?thesis . qed corollary connected_punctured_universe: "2 \ DIM('N::euclidean_space) \ connected(- {a::'N})" by (simp add: path_connected_punctured_universe path_connected_imp_connected) proposition path_connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "path_connected(sphere a r)" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (simp) next case equal then show ?thesis by (simp) next case greater then have eq: "(sphere (0::'a) r) = (\x. (r / norm x) *\<^sub>R x) ` (- {0::'a})" by (force simp: image_iff split: if_split_asm) have "continuous_on (- {0::'a}) (\x. (r / norm x) *\<^sub>R x)" by (intro continuous_intros) auto then have "path_connected ((\x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))" by (intro path_connected_continuous_image path_connected_punctured_universe assms) with eq have "path_connected (sphere (0::'a) r)" by auto then have "path_connected((+) a ` (sphere (0::'a) r))" by (simp add: path_connected_translation) then show ?thesis by (metis add.right_neutral sphere_translation) qed lemma connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "connected(sphere a r)" using path_connected_sphere [OF assms] by (simp add: path_connected_imp_connected) corollary path_connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" and 2: "2 \ DIM('a)" shows "path_connected (- s)" proof (cases "s = {}") case True then show ?thesis using convex_imp_path_connected by auto next case False then obtain a where "a \ s" by auto { fix x y assume "x \ s" "y \ s" then have "x \ a" "y \ a" using \a \ s\ by auto then have bxy: "bounded(insert x (insert y s))" by (simp add: \bounded s\) then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "s \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) define C where "C = B / norm(x - a)" { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + C *\<^sub>R (x - a)) \ s" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" using \x \ a\ \0 \ u\ Bx by (auto simp add: C_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = (1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x" using CC by (simp add: field_simps) also have "\ = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = x + ((1 / (1 + C * u - u)) *\<^sub>R a + ((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))" using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" by (simp add: algebra_simps) have False using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) apply (rule \a \ s\) apply (drule_tac x="a + (1 + (C - 1) * u) *\<^sub>R (x - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (C - 1) * u)" in spec) using \x \ a\ \x \ s\ \0 \ u\ CC apply (auto simp: xeq) done } then have pcx: "path_component (- s) x (a + C *\<^sub>R (x - a))" by (force simp: closed_segment_def intro!: path_component_linepath) define D where "D = B / norm(y - a)" \ \massive duplication with the proof above\ { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \ s" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" using \y \ a\ \0 \ u\ By by (auto simp add: D_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = (1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y" using DD by (simp add: field_simps) also have "\ = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = y + ((1 / (1 + D * u - u)) *\<^sub>R a + ((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))" using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" by (simp add: algebra_simps) have False using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) apply (rule \a \ s\) apply (drule_tac x="a + (1 + (D - 1) * u) *\<^sub>R (y - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (D - 1) * u)" in spec) using \y \ a\ \y \ s\ \0 \ u\ DD apply (auto simp: xeq) done } then have pdy: "path_component (- s) y (a + D *\<^sub>R (y - a))" by (force simp: closed_segment_def intro!: path_component_linepath) have pyx: "path_component (- s) (a + D *\<^sub>R (y - a)) (a + C *\<^sub>R (x - a))" apply (rule path_component_of_subset [of "sphere a B"]) using \s \ ball a B\ apply (force simp: ball_def dist_norm norm_minus_commute) apply (rule path_connected_sphere [OF 2, of a B, simplified path_connected_component, rule_format]) using \x \ a\ using \y \ a\ B apply (auto simp: dist_norm C_def D_def) done have "path_component (- s) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) } then show ?thesis by (auto simp: path_connected_component) qed lemma connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" "2 \ DIM('a)" shows "connected (- s)" using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast lemma connected_diff_ball: fixes s :: "'a :: euclidean_space set" assumes "connected s" "cball a r \ s" "2 \ DIM('a)" shows "connected (s - ball a r)" apply (rule connected_diff_open_from_closed [OF ball_subset_cball]) using assms connected_sphere apply (auto simp: cball_diff_eq_sphere dist_norm) done proposition connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "connected(S - {a::'N})" proof (cases "a \ S") case True with \open S\ obtain \ where "\ > 0" and \: "cball a \ \ S" using open_contains_cball_eq by blast have "dist a (a + \ *\<^sub>R (SOME i. i \ Basis)) = \" by (simp add: dist_norm SOME_Basis \0 < \\ less_imp_le) with \ have "\{S - ball a r |r. 0 < r \ r < \} \ {} \ False" apply (drule_tac c="a + scaleR (\) ((SOME i. i \ Basis))" in subsetD) by auto then have nonemp: "(\{S - ball a r |r. 0 < r \ r < \}) = {} \ False" by auto have con: "\r. r < \ \ connected (S - ball a r)" using \ by (force intro: connected_diff_ball [OF \connected S\ _ 2]) have "x \ \{S - ball a r |r. 0 < r \ r < \}" if "x \ S - {a}" for x apply (rule UnionI [of "S - ball a (min \ (dist a x) / 2)"]) using that \0 < \\ apply simp_all apply (rule_tac x="min \ (dist a x) / 2" in exI) apply auto done then have "S - {a} = \{S - ball a r | r. 0 < r \ r < \}" by auto then show ?thesis by (auto intro: connected_Union con dest!: nonemp) next case False then show ?thesis by (simp add: \connected S\) qed corollary path_connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "path_connected(S - {a::'N})" by (simp add: assms connected_open_delete connected_open_path_connected open_delete) corollary path_connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ path_connected(ball a r - {a::'N})" by (simp add: path_connected_open_delete) corollary connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ connected(ball a r - {a::'N})" by (simp add: connected_open_delete) corollary connected_open_delete_finite: fixes S T::"'a::euclidean_space set" assumes S: "open S" "connected S" and 2: "2 \ DIM('a)" and "finite T" shows "connected(S - T)" using \finite T\ S proof (induct T) case empty show ?case using \connected S\ by simp next case (insert x F) then have "connected (S-F)" by auto moreover have "open (S - F)" using finite_imp_closed[OF \finite F\] \open S\ by auto ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto thus ?case by (metis Diff_insert) qed lemma sphere_1D_doubleton_zero: assumes 1: "DIM('a) = 1" and "r > 0" obtains x y::"'a::euclidean_space" where "sphere 0 r = {x,y} \ dist x y = 2*r" proof - obtain b::'a where b: "Basis = {b}" using 1 card_1_singletonE by blast show ?thesis proof (intro that conjI) have "x = norm x *\<^sub>R b \ x = - norm x *\<^sub>R b" if "r = norm x" for x proof - have xb: "(x \ b) *\<^sub>R b = x" using euclidean_representation [of x, unfolded b] by force then have "norm ((x \ b) *\<^sub>R b) = norm x" by simp with b have "\x \ b\ = norm x" using norm_Basis by (simp add: b) with xb show ?thesis apply (simp add: abs_if split: if_split_asm) apply (metis add.inverse_inverse real_vector.scale_minus_left xb) done qed with \r > 0\ b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}" by (force simp: sphere_def dist_norm) have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)" by (simp add: dist_norm) also have "\ = norm ((2*r) *\<^sub>R b)" by (metis mult_2 scaleR_add_left) also have "\ = 2*r" using \r > 0\ b norm_Basis by fastforce finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" . qed qed lemma sphere_1D_doubleton: fixes a :: "'a :: euclidean_space" assumes "DIM('a) = 1" and "r > 0" obtains x y where "sphere a r = {x,y} \ dist x y = 2*r" proof - have "sphere a r = (+) a ` sphere 0 r" by (metis add.right_neutral sphere_translation) then show ?thesis using sphere_1D_doubleton_zero [OF assms] by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that) qed lemma psubset_sphere_Compl_connected: fixes S :: "'a::euclidean_space set" assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S \ sphere a r" using S by blast obtain b where "dist a b = r" and "b \ S" using S mem_sphere by blast have CS: "- S = {x. dist a x \ r \ (x \ S)} \ {x. r \ dist a x \ (x \ S)}" by auto have "{x. dist a x \ r \ x \ S} \ {x. r \ dist a x \ x \ S} \ {}" using \b \ S\ \dist a b = r\ by blast moreover have "connected {x. dist a x \ r \ x \ S}" apply (rule connected_intermediate_closure [of "ball a r"]) using assms by auto moreover have "connected {x. r \ dist a x \ x \ S}" apply (rule connected_intermediate_closure [of "- cball a r"]) using assms apply (auto intro: connected_complement_bounded_convex) apply (metis ComplI interior_cball interior_closure mem_ball not_less) done ultimately show ?thesis by (simp add: CS connected_Un) qed subsection\Every annulus is a connected set\ lemma path_connected_2DIM_I: fixes a :: "'N::euclidean_space" assumes 2: "2 \ DIM('N)" and pc: "path_connected {r. 0 \ r \ P r}" shows "path_connected {x. P(norm(x - a))}" proof - have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}" by force moreover have "path_connected {x::'N. P(norm x)}" proof - let ?D = "{x. 0 \ x \ P x} \ sphere (0::'N) 1" have "x \ (\z. fst z *\<^sub>R snd z) ` ?D" if "P (norm x)" for x::'N proof (cases "x=0") case True with that show ?thesis apply (simp add: image_iff) apply (rule_tac x=0 in exI, simp) using vector_choose_size zero_le_one by blast next case False with that show ?thesis by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto qed then have *: "{x::'N. P(norm x)} = (\z. fst z *\<^sub>R snd z) ` ?D" by auto have "continuous_on ?D (\z:: real\'N. fst z *\<^sub>R snd z)" by (intro continuous_intros) moreover have "path_connected ?D" by (metis path_connected_Times [OF pc] path_connected_sphere 2) ultimately show ?thesis apply (subst *) apply (rule path_connected_continuous_image, auto) done qed ultimately show ?thesis using path_connected_translation by metis qed proposition path_connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N)" shows "path_connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms]) proposition connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N::euclidean_space)" shows "connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) subsection\<^marker>\tag unimportant\\Relations between components and path components\ lemma open_connected_component: fixes s :: "'a::real_normed_vector set" shows "open s \ open (connected_component_set s x)" apply (simp add: open_contains_ball, clarify) apply (rename_tac y) apply (drule_tac x=y in bspec) apply (simp add: connected_component_in, clarify) apply (rule_tac x=e in exI) by (metis mem_Collect_eq connected_component_eq connected_component_maximal centre_in_ball connected_ball) corollary open_components: fixes s :: "'a::real_normed_vector set" shows "\open u; s \ components u\ \ open s" by (simp add: components_iff) (metis open_connected_component) lemma in_closure_connected_component: fixes s :: "'a::real_normed_vector set" assumes x: "x \ s" and s: "open s" shows "x \ closure (connected_component_set s y) \ x \ connected_component_set s y" proof - { assume "x \ closure (connected_component_set s y)" moreover have "x \ connected_component_set s x" using x by simp ultimately have "x \ connected_component_set s y" using s by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component) } then show ?thesis by (auto simp: closure_def) qed lemma connected_disjoint_Union_open_pick: assumes "pairwise disjnt B" "\S. S \ A \ connected S \ S \ {}" "\S. S \ B \ open S" "\A \ \B" "S \ A" obtains T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" proof - have "S \ \B" "connected S" "S \ {}" using assms \S \ A\ by blast+ then obtain T where "T \ B" "S \ T \ {}" by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute) have 1: "open T" by (simp add: \T \ B\ assms) have 2: "open (\(B-{T}))" using assms by blast have 3: "S \ T \ \(B - {T})" using \S \ \B\ by blast have "T \ \(B - {T}) = {}" using \T \ B\ \pairwise disjnt B\ by (auto simp: pairwise_def disjnt_def) then have 4: "T \ \(B - {T}) \ S = {}" by auto from connectedD [OF \connected S\ 1 2 3 4] have "S \ \(B-{T}) = {}" by (auto simp: Int_commute \S \ T \ {}\) with \T \ B\ have "S \ T" using "3" by auto show ?thesis using \S \ \(B - {T}) = {}\ \S \ T\ \T \ B\ that by auto qed lemma connected_disjoint_Union_open_subset: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A \ B" proof fix S assume "S \ A" obtain T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" apply (rule connected_disjoint_Union_open_pick [OF B, of A]) using SA SB \S \ A\ by auto moreover obtain S' where "S' \ A" "T \ S'" "T \ \(A - {S'}) = {}" apply (rule connected_disjoint_Union_open_pick [OF A, of B]) using SA SB \T \ B\ by auto ultimately have "S' = S" by (metis A Int_subset_iff SA \S \ A\ disjnt_def inf.orderE pairwise_def) with \T \ S'\ have "T \ S" by simp with \S \ T\ have "S = T" by blast with \T \ B\ show "S \ B" by simp qed lemma connected_disjoint_Union_open_unique: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A = B" by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms) proposition components_open_unique: fixes S :: "'a::real_normed_vector set" assumes "pairwise disjnt A" "\A = S" "\X. X \ A \ open X \ connected X \ X \ {}" shows "components S = A" proof - have "open S" using assms by blast show ?thesis apply (rule connected_disjoint_Union_open_unique) apply (simp add: components_eq disjnt_def pairwise_def) using \open S\ apply (simp_all add: assms open_components in_components_connected in_components_nonempty) done qed subsection\<^marker>\tag unimportant\\Existence of unbounded components\ lemma cobounded_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes "bounded (-s)" shows "\x. x \ s \ \ bounded (connected_component_set s x)" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-s \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) have unbounded_inner: "\ bounded {x. inner i x \ B}" apply (auto simp: bounded_def dist_norm) apply (rule_tac x="x + (max B e + 1 + \i \ x\) *\<^sub>R i" in exI) apply simp using i apply (auto simp: algebra_simps) done have **: "{x. B \ i \ x} \ connected_component_set s (B *\<^sub>R i)" apply (rule connected_component_maximal) apply (auto simp: i intro: convex_connected convex_halfspace_ge [of B]) apply (rule *) apply (rule order_trans [OF _ Basis_le_norm [OF i]]) by (simp add: inner_commute) have "B *\<^sub>R i \ s" by (rule *) (simp add: norm_Basis [OF i]) then show ?thesis apply (rule_tac x="B *\<^sub>R i" in exI, clarify) apply (frule bounded_subset [of _ "{x. B \ i \ x}", OF _ **]) using unbounded_inner apply blast done qed lemma cobounded_unique_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes bs: "bounded (-s)" and "2 \ DIM('a)" and bo: "\ bounded(connected_component_set s x)" "\ bounded(connected_component_set s y)" shows "connected_component_set s x = connected_component_set s y" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-s \ ball 0 B" using bounded_subset_ballD [OF bs, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) have ccb: "connected (- ball 0 B :: 'a set)" using assms by (auto intro: connected_complement_bounded_convex) obtain x' where x': "connected_component s x x'" "norm x' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) obtain y' where y': "connected_component s y y'" "norm y' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) have x'y': "connected_component s x' y'" apply (simp add: connected_component_def) apply (rule_tac x="- ball 0 B" in exI) using x' y' apply (auto simp: ccb dist_norm *) done show ?thesis apply (rule connected_component_eq) using x' y' x'y' by (metis (no_types, lifting) connected_component_eq_empty connected_component_eq_eq connected_component_idemp connected_component_in) qed lemma cobounded_unbounded_components: fixes s :: "'a :: euclidean_space set" shows "bounded (-s) \ \c. c \ components s \ \bounded c" by (metis cobounded_unbounded_component components_def imageI) lemma cobounded_unique_unbounded_components: fixes s :: "'a :: euclidean_space set" shows "\bounded (- s); c \ components s; \ bounded c; c' \ components s; \ bounded c'; 2 \ DIM('a)\ \ c' = c" unfolding components_iff by (metis cobounded_unique_unbounded_component) lemma cobounded_has_bounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (- S)" "\ connected S" "2 \ DIM('a)" obtains C where "C \ components S" "bounded C" by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) subsection\The \inside\ and \outside\ of a Set\ text\<^marker>\tag important\\The inside comprises the points in a bounded connected component of the set's complement. The outside comprises the points in unbounded connected component of the complement.\ definition\<^marker>\tag important\ inside where "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" definition\<^marker>\tag important\ outside where "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) lemma inside_no_overlap [simp]: "inside S \ S = {}" by (auto simp: inside_def) lemma outside_no_overlap [simp]: "outside S \ S = {}" by (auto simp: outside_def) lemma inside_Int_outside [simp]: "inside S \ outside S = {}" by (auto simp: inside_def outside_def) lemma inside_Un_outside [simp]: "inside S \ outside S = (- S)" by (auto simp: inside_def outside_def) lemma inside_eq_outside: "inside S = outside S \ S = UNIV" by (auto simp: inside_def outside_def) lemma inside_outside: "inside S = (- (S \ outside S))" by (force simp: inside_def outside) lemma outside_inside: "outside S = (- (S \ inside S))" by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap) lemma union_with_inside: "S \ inside S = - outside S" by (auto simp: inside_outside) (simp add: outside_inside) lemma union_with_outside: "S \ outside S = - inside S" by (simp add: inside_outside) lemma outside_mono: "S \ T \ outside T \ outside S" by (auto simp: outside bounded_subset connected_component_mono) lemma inside_mono: "S \ T \ inside S - T \ inside T" by (auto simp: inside_def bounded_subset connected_component_mono) lemma segment_bound_lemma: fixes u::real assumes "x \ B" "y \ B" "0 \ u" "u \ 1" shows "(1 - u) * x + u * y \ B" proof - obtain dx dy where "dx \ 0" "dy \ 0" "x = B + dx" "y = B + dy" using assms by auto (metis add.commute diff_add_cancel) with \0 \ u\ \u \ 1\ show ?thesis by (simp add: add_increasing2 mult_left_le field_simps) qed lemma cobounded_outside: fixes S :: "'a :: real_normed_vector set" assumes "bounded S" shows "bounded (- outside S)" proof - obtain B where B: "B>0" "S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto { fix x::'a and C::real assume Bno: "B \ norm x" and C: "0 < C" have "\y. connected_component (- S) x y \ norm y > C" proof (cases "x = 0") case True with B Bno show ?thesis by force next case False have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \ - ball 0 B" proof fix w assume "w \ closed_segment x (((B + C) / norm x) *\<^sub>R x)" then obtain u where w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \ u" "u \ 1" by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric]) with False B C have "B \ (1 - u) * norm x + u * (B + C)" using segment_bound_lemma [of B "norm x" "B + C" u] Bno by simp with False B C show "w \ - ball 0 B" using distrib_right [of _ _ "norm x"] by (simp add: ball_def w not_less) qed also have "... \ -S" by (simp add: B) finally have "\T. connected T \ T \ - S \ x \ T \ ((B + C) / norm x) *\<^sub>R x \ T" by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp with False B show ?thesis by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def) qed } then show ?thesis apply (simp add: outside_def assms) apply (rule bounded_subset [OF bounded_ball [of 0 B]]) apply (force simp: dist_norm not_less bounded_pos) done qed lemma unbounded_outside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ \ bounded(outside S)" using cobounded_imp_unbounded cobounded_outside by blast lemma bounded_inside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ bounded(inside S)" by (simp add: bounded_Int cobounded_outside inside_outside) lemma connected_outside: fixes S :: "'a::euclidean_space set" assumes "bounded S" "2 \ DIM('a)" shows "connected(outside S)" apply (clarsimp simp add: connected_iff_connected_component outside) apply (rule_tac s="connected_component_set (- S) x" in connected_component_of_subset) apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq) apply clarify apply (metis connected_component_eq_eq connected_component_in) done lemma outside_connected_component_lt: "outside S = {x. \B. \y. B < norm(y) \ connected_component (- S) x y}" apply (auto simp: outside bounded_def dist_norm) apply (metis diff_0 norm_minus_cancel not_less) by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6)) lemma outside_connected_component_le: "outside S = {x. \B. \y. B \ norm(y) \ connected_component (- S) x y}" apply (simp add: outside_connected_component_lt) apply (simp add: Set.set_eq_iff) by (meson gt_ex leD le_less_linear less_imp_le order.trans) lemma not_outside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" and "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B < norm(y) \ \ connected_component (- S) x y}" proof - obtain B::real where B: "0 < B" and Bno: "\x. x \ S \ norm x \ B" using S [simplified bounded_pos] by auto { fix y::'a and z::'a assume yz: "B < norm z" "B < norm y" have "connected_component (- cball 0 B) y z" apply (rule connected_componentI [OF _ subset_refl]) apply (rule connected_complement_bounded_convex) using assms yz by (auto simp: dist_norm) then have "connected_component (- S) y z" apply (rule connected_component_of_subset) apply (metis Bno Compl_anti_mono mem_cball_0 subset_iff) done } note cyz = this show ?thesis apply (auto simp: outside) apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le) apply (simp add: bounded_pos) by (metis B connected_component_trans cyz not_le) qed lemma not_outside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B \ norm(y) \ \ connected_component (- S) x y}" apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms]) by (meson gt_ex less_le_trans) lemma inside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) lemma inside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) lemma inside_subset: assumes "connected U" and "\ bounded U" and "T \ U = - S" shows "inside S \ T" apply (auto simp: inside_def) by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal Compl_iff Un_iff assms subsetI) lemma frontier_not_empty: fixes S :: "'a :: real_normed_vector set" shows "\S \ {}; S \ UNIV\ \ frontier S \ {}" using connected_Int_frontier [of UNIV S] by auto lemma frontier_eq_empty: fixes S :: "'a :: real_normed_vector set" shows "frontier S = {} \ S = {} \ S = UNIV" using frontier_UNIV frontier_empty frontier_not_empty by blast lemma frontier_of_connected_component_subset: fixes S :: "'a::real_normed_vector set" shows "frontier(connected_component_set S x) \ frontier S" proof - { fix y assume y1: "y \ closure (connected_component_set S x)" and y2: "y \ interior (connected_component_set S x)" have "y \ closure S" using y1 closure_mono connected_component_subset by blast moreover have "z \ interior (connected_component_set S x)" if "0 < e" "ball y e \ interior S" "dist y z < e" for e z proof - have "ball y e \ connected_component_set S y" apply (rule connected_component_maximal) using that interior_subset mem_ball apply auto done then show ?thesis using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \0 < e\ y2) qed then have "y \ interior S" using y2 by (force simp: open_contains_ball_eq [OF open_interior]) ultimately have "y \ frontier S" by (auto simp: frontier_def) } then show ?thesis by (auto simp: frontier_def) qed lemma frontier_Union_subset_closure: fixes F :: "'a::real_normed_vector set set" shows "frontier(\F) \ closure(\t \ F. frontier t)" proof - have "\y\F. \y\frontier y. dist y x < e" if "T \ F" "y \ T" "dist y x < e" "x \ interior (\F)" "0 < e" for x y e T proof (cases "x \ T") case True with that show ?thesis by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono) next case False have 1: "closed_segment x y \ T \ {}" using \y \ T\ by blast have 2: "closed_segment x y - T \ {}" using False by blast obtain c where "c \ closed_segment x y" "c \ frontier T" using False connected_Int_frontier [OF connected_segment 1 2] by auto then show ?thesis proof - have "norm (y - x) < e" by (metis dist_norm \dist y x < e\) moreover have "norm (c - x) \ norm (y - x)" by (simp add: \c \ closed_segment x y\ segment_bound(1)) ultimately have "norm (c - x) < e" by linarith then show ?thesis by (metis (no_types) \c \ frontier T\ dist_norm that(1)) qed qed then show ?thesis by (fastforce simp add: frontier_def closure_approachable) qed lemma frontier_Union_subset: fixes F :: "'a::real_normed_vector set set" shows "finite F \ frontier(\F) \ (\t \ F. frontier t)" by (rule order_trans [OF frontier_Union_subset_closure]) (auto simp: closure_subset_eq) lemma frontier_of_components_subset: fixes S :: "'a::real_normed_vector set" shows "C \ components S \ frontier C \ frontier S" by (metis Path_Connected.frontier_of_connected_component_subset components_iff) lemma frontier_of_components_closed_complement: fixes S :: "'a::real_normed_vector set" shows "\closed S; C \ components (- S)\ \ frontier C \ S" using frontier_complement frontier_of_components_subset frontier_subset_eq by blast lemma frontier_minimal_separating_closed: fixes S :: "'a::real_normed_vector set" assumes "closed S" and nconn: "\ connected(- S)" and C: "C \ components (- S)" and conn: "\T. \closed T; T \ S\ \ connected(- T)" shows "frontier C = S" proof (rule ccontr) assume "frontier C \ S" then have "frontier C \ S" using frontier_of_components_closed_complement [OF \closed S\ C] by blast then have "connected(- (frontier C))" by (simp add: conn) have "\ connected(- (frontier C))" unfolding connected_def not_not proof (intro exI conjI) show "open C" using C \closed S\ open_components by blast show "open (- closure C)" by blast show "C \ - closure C \ - frontier C = {}" using closure_subset by blast show "C \ - frontier C \ {}" using C \open C\ components_eq frontier_disjoint_eq by fastforce show "- frontier C \ C \ - closure C" by (simp add: \open C\ closed_Compl frontier_closures) then show "- closure C \ - frontier C \ {}" by (metis (no_types, lifting) C Compl_subset_Compl_iff \frontier C \ S\ compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb) qed then show False using \connected (- frontier C)\ by blast qed lemma connected_component_UNIV [simp]: fixes x :: "'a::real_normed_vector" shows "connected_component_set UNIV x = UNIV" using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV by auto lemma connected_component_eq_UNIV: fixes x :: "'a::real_normed_vector" shows "connected_component_set s x = UNIV \ s = UNIV" using connected_component_in connected_component_UNIV by blast lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" by (auto simp: components_eq_sing_iff) lemma interior_inside_frontier: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "interior s \ inside (frontier s)" proof - { fix x y assume x: "x \ interior s" and y: "y \ s" and cc: "connected_component (- frontier s) x y" have "connected_component_set (- frontier s) x \ frontier s \ {}" apply (rule connected_Int_frontier, simp) apply (metis IntI cc connected_component_in connected_component_refl empty_iff interiorE mem_Collect_eq rev_subsetD x) using y cc by blast then have "bounded (connected_component_set (- frontier s) x)" using connected_component_in by auto } then show ?thesis apply (auto simp: inside_def frontier_def) apply (rule classical) apply (rule bounded_subset [OF assms], blast) done qed lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)" by (simp add: inside_def) lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)" using inside_empty inside_Un_outside by blast lemma inside_same_component: "\connected_component (- s) x y; x \ inside s\ \ y \ inside s" using connected_component_eq connected_component_in by (fastforce simp add: inside_def) lemma outside_same_component: "\connected_component (- s) x y; x \ outside s\ \ y \ outside s" using connected_component_eq connected_component_in by (fastforce simp add: outside_def) lemma convex_in_outside: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes s: "convex s" and z: "z \ s" shows "z \ outside s" proof (cases "s={}") case True then show ?thesis by simp next case False then obtain a where "a \ s" by blast with z have zna: "z \ a" by auto { assume "bounded (connected_component_set (- s) z)" with bounded_pos_less obtain B where "B>0" and B: "\x. connected_component (- s) z x \ norm x < B" by (metis mem_Collect_eq) define C where "C = (B + 1 + norm z) / norm (z-a)" have "C > 0" using \0 < B\ zna by (simp add: C_def field_split_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" by (metis add_diff_cancel norm_triangle_ineq3) moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" using zna \B>0\ by (simp add: C_def le_max_iff_disj) ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith { fix u::real assume u: "0\u" "u\1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \ s" then have Cpos: "1 + u * C > 0" by (meson \0 < C\ add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" by (simp add: scaleR_add_left [symmetric] field_split_simps) then have False using convexD_alt [OF s \a \ s\ ins, of "1/(u*C + 1)"] \C>0\ \z \ s\ Cpos u by (simp add: * field_split_simps) } note contra = this have "connected_component (- s) z (z + C *\<^sub>R (z-a))" apply (rule connected_componentI [OF connected_segment [of z "z + C *\<^sub>R (z-a)"]]) apply (simp add: closed_segment_def) using contra apply auto done then have False using zna B [of "z + C *\<^sub>R (z-a)"] C by (auto simp: field_split_simps max_mult_distrib_right) } then show ?thesis by (auto simp: outside_def z) qed lemma outside_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "convex s" shows "outside s = - s" by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2) lemma outside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "outside {x} = -{x}" by (auto simp: outside_convex) lemma inside_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "convex s \ inside s = {}" by (simp add: inside_outside outside_convex) lemma inside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "inside {x} = {}" by (auto simp: inside_convex) lemma outside_subset_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "\convex t; s \ t\ \ - t \ outside s" using outside_convex outside_mono by blast lemma outside_Un_outside_Un: fixes S :: "'a::real_normed_vector set" assumes "S \ outside(T \ U) = {}" shows "outside(T \ U) \ outside(T \ S)" proof fix x assume x: "x \ outside (T \ U)" have "Y \ - S" if "connected Y" "Y \ - T" "Y \ - U" "x \ Y" "u \ Y" for u Y proof - have "Y \ connected_component_set (- (T \ U)) x" by (simp add: connected_component_maximal that) also have "\ \ outside(T \ U)" by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) finally have "Y \ outside(T \ U)" . with assms show ?thesis by auto qed with x show "x \ outside (T \ S)" by (simp add: outside_connected_component_lt connected_component_def) meson qed lemma outside_frontier_misses_closure: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "outside(frontier s) \ - closure s" unfolding outside_inside Lattices.boolean_algebra_class.compl_le_compl_iff proof - { assume "interior s \ inside (frontier s)" hence "interior s \ inside (frontier s) = inside (frontier s)" by (simp add: subset_Un_eq) then have "closure s \ frontier s \ inside (frontier s)" using frontier_def by auto } then show "closure s \ frontier s \ inside (frontier s)" using interior_inside_frontier [OF assms] by blast qed lemma outside_frontier_eq_complement_closure: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded s" "convex s" shows "outside(frontier s) = - closure s" by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure outside_subset_convex subset_antisym) lemma inside_frontier_eq_interior: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "\bounded s; convex s\ \ inside(frontier s) = interior s" apply (simp add: inside_outside outside_frontier_eq_complement_closure) using closure_subset interior_subset apply (auto simp: frontier_def) done lemma open_inside: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "open (inside s)" proof - { fix x assume x: "x \ inside s" have "open (connected_component_set (- s) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- s) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x) then have "\e>0. ball x e \ inside s" by (metis e dist_commute inside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma open_outside: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "open (outside s)" proof - { fix x assume x: "x \ outside s" have "open (connected_component_set (- s) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- s) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis Int_iff outside_def connected_component_refl_eq x) then have "\e>0. ball x e \ outside s" by (metis e dist_commute outside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma closure_inside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "closure(inside s) \ s \ inside s" by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside) lemma frontier_inside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "frontier(inside s) \ s" proof - have "closure (inside s) \ - inside s = closure (inside s) - interior (inside s)" by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside) moreover have "- inside s \ - outside s = s" by (metis (no_types) compl_sup double_compl inside_Un_outside) moreover have "closure (inside s) \ - outside s" by (metis (no_types) assms closure_inside_subset union_with_inside) ultimately have "closure (inside s) - interior (inside s) \ s" by blast then show ?thesis by (simp add: frontier_def open_inside interior_open) qed lemma closure_outside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "closure(outside s) \ s \ outside s" apply (rule closure_minimal, simp) by (metis assms closed_open inside_outside open_inside) lemma frontier_outside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "frontier(outside s) \ s" apply (simp add: frontier_def open_outside interior_open) by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup.commute) lemma inside_complement_unbounded_connected_empty: "\connected (- s); \ bounded (- s)\ \ inside s = {}" apply (simp add: inside_def) by (meson Compl_iff bounded_subset connected_component_maximal order_refl) lemma inside_bounded_complement_connected_empty: fixes s :: "'a::{real_normed_vector, perfect_space} set" shows "\connected (- s); bounded s\ \ inside s = {}" by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded) lemma inside_inside: assumes "s \ inside t" shows "inside s - t \ inside t" unfolding inside_def proof clarify fix x assume x: "x \ t" "x \ s" and bo: "bounded (connected_component_set (- s) x)" show "bounded (connected_component_set (- t) x)" proof (cases "s \ connected_component_set (- t) x = {}") case True show ?thesis apply (rule bounded_subset [OF bo]) apply (rule connected_component_maximal) using x True apply auto done next case False then show ?thesis using assms [unfolded inside_def] x apply (simp add: disjoint_iff_not_equal, clarify) apply (drule subsetD, assumption, auto) by (metis (no_types, hide_lams) ComplI connected_component_eq_eq) qed qed lemma inside_inside_subset: "inside(inside s) \ s" using inside_inside union_with_outside by fastforce lemma inside_outside_intersect_connected: "\connected t; inside s \ t \ {}; outside s \ t \ {}\ \ s \ t \ {}" apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify) by (metis (no_types, hide_lams) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl) lemma outside_bounded_nonempty: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded s" shows "outside s \ {}" by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball double_complement order_refl outside_convex outside_def) lemma outside_compact_in_open: fixes s :: "'a :: {real_normed_vector,perfect_space} set" assumes s: "compact s" and t: "open t" and "s \ t" "t \ {}" shows "outside s \ t \ {}" proof - have "outside s \ {}" by (simp add: compact_imp_bounded outside_bounded_nonempty s) with assms obtain a b where a: "a \ outside s" and b: "b \ t" by auto show ?thesis proof (cases "a \ t") case True with a show ?thesis by blast next case False have front: "frontier t \ - s" using \s \ t\ frontier_disjoint_eq t by auto { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- t)" and pf: "pathfinish \ \ frontier t" and ps: "pathstart \ = a" define c where "c = pathfinish \" have "c \ -s" unfolding c_def using front pf by blast moreover have "open (-s)" using s compact_imp_closed by blast ultimately obtain \::real where "\ > 0" and \: "cball c \ \ -s" using open_contains_cball[of "-s"] s by blast then obtain d where "d \ t" and d: "dist d c < \" using closure_approachable [of c t] pf unfolding c_def by (metis Diff_iff frontier_def) then have "d \ -s" using \ using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq) have pimg_sbs_cos: "path_image \ \ -s" using pimg_sbs apply (auto simp: path_image_def) apply (drule subsetD) using \c \ - s\ \s \ t\ interior_subset apply (auto simp: c_def) done have "closed_segment c d \ cball c \" apply (simp add: segment_convex_hull) apply (rule hull_minimal) using \\ > 0\ d apply (auto simp: dist_commute) done with \ have "closed_segment c d \ -s" by blast moreover have con_gcd: "connected (path_image \ \ closed_segment c d)" by (rule connected_Un) (auto simp: c_def \path \\ connected_path_image) ultimately have "connected_component (- s) a d" unfolding connected_component_def using pimg_sbs_cos ps by blast then have "outside s \ t \ {}" using outside_same_component [OF _ a] by (metis IntI \d \ t\ empty_iff) } note * = this have pal: "pathstart (linepath a b) \ closure (- t)" by (auto simp: False closure_def) show ?thesis by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b) qed qed lemma inside_inside_compact_connected: fixes s :: "'a :: euclidean_space set" assumes s: "closed s" and t: "compact t" and "connected t" "s \ inside t" shows "inside s \ inside t" proof (cases "inside t = {}") case True with assms show ?thesis by auto next case False consider "DIM('a) = 1" | "DIM('a) \ 2" using antisym not_less_eq_eq by fastforce then show ?thesis proof cases case 1 then show ?thesis using connected_convex_1_gen assms False inside_convex by blast next case 2 have coms: "compact s" using assms apply (simp add: s compact_eq_bounded_closed) by (meson bounded_inside bounded_subset compact_imp_bounded) then have bst: "bounded (s \ t)" by (simp add: compact_imp_bounded t) then obtain r where "0 < r" and r: "s \ t \ ball 0 r" using bounded_subset_ballD by blast have outst: "outside s \ outside t \ {}" proof - have "- ball 0 r \ outside s" apply (rule outside_subset_convex) using r by auto moreover have "- ball 0 r \ outside t" apply (rule outside_subset_convex) using r by auto ultimately show ?thesis by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap) qed have "s \ t = {}" using assms by (metis disjoint_iff_not_equal inside_no_overlap subsetCE) moreover have "outside s \ inside t \ {}" by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open t) ultimately have "inside s \ t = {}" using inside_outside_intersect_connected [OF \connected t\, of s] by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) then show ?thesis using inside_inside [OF \s \ inside t\] by blast qed qed lemma connected_with_inside: fixes s :: "'a :: real_normed_vector set" assumes s: "closed s" and cons: "connected s" shows "connected(s \ inside s)" proof (cases "s \ inside s = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ s" "b \ inside s" by blast have *: "\y t. y \ s \ connected t \ a \ t \ y \ t \ t \ (s \ inside s)" if "a \ (s \ inside s)" for a using that proof assume "a \ s" then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x="{a}" in exI, simp) done next assume a: "a \ inside s" show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside s"]) using a apply (simp add: closure_def) apply (simp add: b) apply (rule_tac x="pathfinish h" in exI) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) using frontier_inside_subset s apply fastforce by (metis (no_types, lifting) frontier_inside_subset insertE insert_Diff interior_eq open_inside pathfinish_in_path_image s subsetCE) qed show ?thesis apply (simp add: connected_iff_connected_component) apply (simp add: connected_component_def) apply (clarify dest!: *) apply (rename_tac u u' t t') apply (rule_tac x="(s \ t \ t')" in exI) apply (auto simp: intro!: connected_Un cons) done qed text\The proof is virtually the same as that above.\ lemma connected_with_outside: fixes s :: "'a :: real_normed_vector set" assumes s: "closed s" and cons: "connected s" shows "connected(s \ outside s)" proof (cases "s \ outside s = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ s" "b \ outside s" by blast have *: "\y t. y \ s \ connected t \ a \ t \ y \ t \ t \ (s \ outside s)" if "a \ (s \ outside s)" for a using that proof assume "a \ s" then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x="{a}" in exI, simp) done next assume a: "a \ outside s" show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside s"]) using a apply (simp add: closure_def) apply (simp add: b) apply (rule_tac x="pathfinish h" in exI) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) using frontier_outside_subset s apply fastforce by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image s subsetCE) qed show ?thesis apply (simp add: connected_iff_connected_component) apply (simp add: connected_component_def) apply (clarify dest!: *) apply (rename_tac u u' t t') apply (rule_tac x="(s \ t \ t')" in exI) apply (auto simp: intro!: connected_Un cons) done qed lemma inside_inside_eq_empty [simp]: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes s: "closed s" and cons: "connected s" shows "inside (inside s) = {}" by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un inside_complement_unbounded_connected_empty unbounded_outside union_with_outside) lemma inside_in_components: "inside s \ components (- s) \ connected(inside s) \ inside s \ {}" apply (simp add: in_components_maximal) apply (auto intro: inside_same_component connected_componentI) apply (metis IntI empty_iff inside_no_overlap) done text\The proof is virtually the same as that above.\ lemma outside_in_components: "outside s \ components (- s) \ connected(outside s) \ outside s \ {}" apply (simp add: in_components_maximal) apply (auto intro: outside_same_component connected_componentI) apply (metis IntI empty_iff outside_no_overlap) done lemma bounded_unique_outside: fixes s :: "'a :: euclidean_space set" shows "\bounded s; DIM('a) \ 2\ \ (c \ components (- s) \ \ bounded c \ c = outside s)" apply (rule iffI) apply (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) by (simp add: connected_outside outside_bounded_nonempty outside_in_components unbounded_outside) subsection\Condition for an open map's image to contain a ball\ proposition ball_subset_open_map_image: fixes f :: "'a::heine_borel \ 'b :: {real_normed_vector,heine_borel}" assumes contf: "continuous_on (closure S) f" and oint: "open (f ` interior S)" and le_no: "\z. z \ frontier S \ r \ norm(f z - f a)" and "bounded S" "a \ S" "0 < r" shows "ball (f a) r \ f ` S" proof (cases "f ` S = UNIV") case True then show ?thesis by simp next case False obtain w where w: "w \ frontier (f ` S)" and dw_le: "\y. y \ frontier (f ` S) \ norm (f a - w) \ norm (f a - y)" apply (rule distance_attains_inf [of "frontier(f ` S)" "f a"]) using \a \ S\ by (auto simp: frontier_eq_empty dist_norm False) then obtain \ where \: "\n. \ n \ f ` S" and tendsw: "\ \ w" by (metis Diff_iff frontier_def closure_sequential) then have "\n. \x \ S. \ n = f x" by force then obtain z where zs: "\n. z n \ S" and fz: "\n. \ n = f (z n)" by metis then obtain y K where y: "y \ closure S" and "strict_mono (K :: nat \ nat)" and Klim: "(z \ K) \ y" using \bounded S\ apply (simp add: compact_closure [symmetric] compact_def) apply (drule_tac x=z in spec) using closure_subset apply force done then have ftendsw: "((\n. f (z n)) \ K) \ w" by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw) have zKs: "\n. (z \ K) n \ S" by (simp add: zs) have fz: "f \ z = \" "(\n. f (z n)) = \" using fz by auto then have "(\ \ K) \ f y" by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially) with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto have rle: "r \ norm (f y - f a)" apply (rule le_no) using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) have **: "(b \ (- S) \ {} \ b - (- S) \ {} \ b \ f \ {}) \ (b \ S \ {}) \ b \ f = {} \ b \ S" for b f and S :: "'b set" by blast show ?thesis apply (rule **) (*such a horrible mess*) apply (rule connected_Int_frontier [where t = "f`S", OF connected_ball]) using \a \ S\ \0 < r\ apply (auto simp: disjoint_iff_not_equal dist_norm) by (metis dw_le norm_minus_commute not_less order_trans rle wy) qed subsubsection\Special characterizations of classes of functions into and out of R.\ proposition embedding_map_into_euclideanreal: assumes "path_connected_space X" shows "embedding_map X euclideanreal f \ continuous_map X euclideanreal f \ inj_on f (topspace X)" proof safe show "continuous_map X euclideanreal f" if "embedding_map X euclideanreal f" using continuous_map_in_subtopology homeomorphic_imp_continuous_map that unfolding embedding_map_def by blast show "inj_on f (topspace X)" if "embedding_map X euclideanreal f" using that homeomorphic_imp_injective_map unfolding embedding_map_def by blast show "embedding_map X euclideanreal f" if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)" proof - obtain g where gf: "\x. x \ topspace X \ g (f x) = x" using inv_into_f_f [OF inj] by auto show ?thesis unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def proof (intro exI conjI) show "continuous_map X (top_of_set (f ` topspace X)) f" by (simp add: cont continuous_map_in_subtopology) let ?S = "f ` topspace X" have eq: "{x \ ?S. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (auto simp: gf) have 1: "g ` ?S \ topspace X" using eq by blast have "openin (top_of_set ?S) {x \ ?S. g x \ T}" if "openin X T" for T proof - have "T \ topspace X" by (simp add: openin_subset that) have RR: "\x \ ?S \ g -` T. \d>0. \x' \ ?S \ ball x d. g x' \ T" proof (clarsimp simp add: gf) have pcS: "path_connectedin euclidean ?S" using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast show "\d>0. \x'\f ` topspace X \ ball (f x) d. g x' \ T" if "x \ T" for x proof - have x: "x \ topspace X" using \T \ topspace X\ \x \ T\ by blast obtain u v d where "0 < d" "u \ topspace X" "v \ topspace X" and sub_fuv: "?S \ {f x - d .. f x + d} \ {f u..f v}" proof (cases "\u \ topspace X. f u < f x") case True then obtain u where u: "u \ topspace X" "f u < f x" .. show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "min (f x - f u) (f v - f x)" show "0 < ?d" by (simp add: \f u < f x\ \f x < f v\) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f v}" by fastforce qed (auto simp: u v) next case False show ?thesis proof let ?d = "f x - f u" show "0 < ?d" by (simp add: u) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f x}" using x u False by auto qed (auto simp: x u) qed next case False note no_u = False show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "f v - f x" show "0 < ?d" by (simp add: v) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f x..f v}" using False by auto qed (auto simp: x v) next case False show ?thesis proof show "f ` topspace X \ {f x - 1..f x + 1} \ {f x..f x}" using False no_u by fastforce qed (auto simp: x) qed qed then obtain h where "pathin X h" "h 0 = u" "h 1 = v" using assms unfolding path_connected_space_def by blast obtain C where "compactin X C" "connectedin X C" "u \ C" "v \ C" proof show "compactin X (h ` {0..1})" using that by (simp add: \pathin X h\ compactin_path_image) show "connectedin X (h ` {0..1})" using \pathin X h\ connectedin_path_image by blast qed (use \h 0 = u\ \h 1 = v\ in auto) have "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) (subtopology X C) g" proof (rule continuous_inverse_map) show "compact_space (subtopology X C)" using \compactin X C\ compactin_subspace by blast show "continuous_map (subtopology X C) euclideanreal f" by (simp add: cont continuous_map_from_subtopology) have "{f u .. f v} \ f ` topspace (subtopology X C)" proof (rule connected_contains_Icc) show "connected (f ` topspace (subtopology X C))" using connectedin_continuous_map_image [OF cont] by (simp add: \compactin X C\ \connectedin X C\ compactin_subset_topspace inf_absorb2) show "f u \ f ` topspace (subtopology X C)" by (simp add: \u \ C\ \u \ topspace X\) show "f v \ f ` topspace (subtopology X C)" by (simp add: \v \ C\ \v \ topspace X\) qed then show "f ` topspace X \ {f x - d..f x + d} \ f ` topspace (subtopology X C)" using sub_fuv by blast qed (auto simp: gf) then have contg: "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) X g" using continuous_map_in_subtopology by blast have "\e>0. \x \ ?S \ {f x - d .. f x + d} \ ball (f x) e. g x \ T" using openin_continuous_map_preimage [OF contg \openin X T\] x \x \ T\ \0 < d\ unfolding openin_euclidean_subtopology_iff by (force simp: gf dist_commute) then obtain e where "e > 0 \ (\x\f ` topspace X \ {f x - d..f x + d} \ ball (f x) e. g x \ T)" by metis with \0 < d\ have "min d e > 0" "\u. u \ topspace X \ \f x - f u\ < min d e \ u \ T" using dist_real_def gf by force+ then show ?thesis by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf) qed qed then obtain d where d: "\r. r \ ?S \ g -` T \ d r > 0 \ (\x \ ?S \ ball r (d r). g x \ T)" by metis show ?thesis unfolding openin_subtopology proof (intro exI conjI) show "{x \ ?S. g x \ T} = (\r \ ?S \ g -` T. ball r (d r)) \ f ` topspace X" using d by (auto simp: gf) qed auto qed then show "continuous_map (top_of_set ?S) X g" by (simp add: continuous_map_def gf) qed (auto simp: gf) qed qed subsubsection \An injective function into R is a homeomorphism and so an open map.\ lemma injective_into_1d_eq_homeomorphism: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on S f" and S: "path_connected S" shows "inj_on f S \ (\g. homeomorphism S (f ` S) f g)" proof show "\g. homeomorphism S (f ` S) f g" if "inj_on f S" proof - have "embedding_map (top_of_set S) euclideanreal f" using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto then show ?thesis by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that) qed qed (metis homeomorphism_def inj_onI) lemma injective_into_1d_imp_open_map: fixes f :: "'a::topological_space \ real" assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T" shows "openin (subtopology euclidean (f ` S)) (f ` T)" using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast lemma homeomorphism_into_1d: fixes f :: "'a::topological_space \ real" assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S" shows "\g. homeomorphism S T f g" using assms injective_into_1d_eq_homeomorphism by blast + end diff --git a/src/HOL/Analysis/Retracts.thy b/src/HOL/Analysis/Retracts.thy --- a/src/HOL/Analysis/Retracts.thy +++ b/src/HOL/Analysis/Retracts.thy @@ -1,2594 +1,2641 @@ section \Absolute Retracts, Absolute Neighbourhood Retracts and Euclidean Neighbourhood Retracts\ theory Retracts imports Brouwer_Fixpoint Continuous_Extension begin text \Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also Euclidean neighbourhood retracts (ENR). We define AR and ANR by specializing the standard definitions for a set to embedding in spaces of higher dimension. John Harrison writes: "This turns out to be sufficient (since any set in \\\<^sup>n\ can be embedded as a closed subset of a convex subset of \\\<^sup>n\<^sup>+\<^sup>1\) to derive the usual definitions, but we need to split them into two implications because of the lack of type quantifiers. Then ENR turns out to be equivalent to ANR plus local compactness."\ definition\<^marker>\tag important\ AR :: "'a::topological_space set \ bool" where "AR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (top_of_set U) S' \ S' retract_of U" definition\<^marker>\tag important\ ANR :: "'a::topological_space set \ bool" where "ANR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (top_of_set U) S' \ (\T. openin (top_of_set U) T \ S' retract_of T)" definition\<^marker>\tag important\ ENR :: "'a::topological_space set \ bool" where "ENR S \ \U. open U \ S retract_of U" text \First, show that we do indeed get the "usual" properties of ARs and ANRs.\ lemma AR_imp_absolute_extensor: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "AR S" and contf: "continuous_on T f" and "f ` T \ S" and cloUT: "closedin (top_of_set U) T" obtains g where "continuous_on U g" "g ` U \ S" "\x. x \ T \ g x = f x" proof - have "aff_dim S < int (DIM('b \ real))" using aff_dim_le_DIM [of S] by simp then obtain C and S' :: "('b * real) set" where C: "convex C" "C \ {}" and cloCS: "closedin (top_of_set C) S'" and hom: "S homeomorphic S'" by (metis that homeomorphic_closedin_convex) then have "S' retract_of C" using \AR S\ by (simp add: AR_def) then obtain r where "S' \ C" and contr: "continuous_on C r" and "r ` C \ S'" and rid: "\x. x\S' \ r x = x" by (auto simp: retraction_def retract_of_def) obtain g h where "homeomorphism S S' g h" using hom by (force simp: homeomorphic_def) then have "continuous_on (f ` T) g" by (meson \f ` T \ S\ continuous_on_subset homeomorphism_def) then have contgf: "continuous_on T (g \ f)" by (metis continuous_on_compose contf) have gfTC: "(g \ f) ` T \ C" proof - have "g ` S = S'" by (metis (no_types) \homeomorphism S S' g h\ homeomorphism_def) with \S' \ C\ \f ` T \ S\ show ?thesis by force qed obtain f' where f': "continuous_on U f'" "f' ` U \ C" "\x. x \ T \ f' x = (g \ f) x" by (metis Dugundji [OF C cloUT contgf gfTC]) show ?thesis proof (rule_tac g = "h \ r \ f'" in that) show "continuous_on U (h \ r \ f')" apply (intro continuous_on_compose f') using continuous_on_subset contr f' apply blast by (meson \homeomorphism S S' g h\ \r ` C \ S'\ continuous_on_subset \f' ` U \ C\ homeomorphism_def image_mono) show "(h \ r \ f') ` U \ S" using \homeomorphism S S' g h\ \r ` C \ S'\ \f' ` U \ C\ by (fastforce simp: homeomorphism_def) show "\x. x \ T \ (h \ r \ f') x = f x" using \homeomorphism S S' g h\ \f ` T \ S\ f' by (auto simp: rid homeomorphism_def) qed qed lemma AR_imp_absolute_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "AR S" "S homeomorphic S'" and clo: "closedin (top_of_set U) S'" shows "S' retract_of U" proof - obtain g h where hom: "homeomorphism S S' g h" using assms by (force simp: homeomorphic_def) have h: "continuous_on S' h" " h ` S' \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done obtain h' where h': "continuous_on U h'" "h' ` U \ S" and h'h: "\x. x \ S' \ h' x = h x" by (blast intro: AR_imp_absolute_extensor [OF \AR S\ h clo]) have [simp]: "S' \ U" using clo closedin_limpt by blast show ?thesis proof (simp add: retraction_def retract_of_def, intro exI conjI) show "continuous_on U (g \ h')" apply (intro continuous_on_compose h') apply (meson hom continuous_on_subset h' homeomorphism_cont1) done show "(g \ h') ` U \ S'" using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) show "\x\S'. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed qed lemma AR_imp_absolute_retract_UNIV: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "AR S" and hom: "S homeomorphic S'" and clo: "closed S'" shows "S' retract_of UNIV" apply (rule AR_imp_absolute_retract [OF \AR S\ hom]) using clo closed_closedin by auto lemma absolute_extensor_imp_AR: fixes S :: "'a::euclidean_space set" assumes "\f :: 'a * real \ 'a. \U T. \continuous_on T f; f ` T \ S; closedin (top_of_set U) T\ \ \g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)" shows "AR S" proof (clarsimp simp: AR_def) fix U and T :: "('a * real) set" assume "S homeomorphic T" and clo: "closedin (top_of_set U) T" then obtain g h where hom: "homeomorphism S T g h" by (force simp: homeomorphic_def) have h: "continuous_on T h" " h ` T \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done obtain h' where h': "continuous_on U h'" "h' ` U \ S" and h'h: "\x\T. h' x = h x" using assms [OF h clo] by blast have [simp]: "T \ U" using clo closedin_imp_subset by auto show "T retract_of U" proof (simp add: retraction_def retract_of_def, intro exI conjI) show "continuous_on U (g \ h')" apply (intro continuous_on_compose h') apply (meson hom continuous_on_subset h' homeomorphism_cont1) done show "(g \ h') ` U \ T" using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) show "\x\T. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed qed lemma AR_eq_absolute_extensor: fixes S :: "'a::euclidean_space set" shows "AR S \ (\f :: 'a * real \ 'a. \U T. continuous_on T f \ f ` T \ S \ closedin (top_of_set U) T \ (\g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)))" apply (rule iffI) apply (metis AR_imp_absolute_extensor) apply (simp add: absolute_extensor_imp_AR) done lemma AR_imp_retract: fixes S :: "'a::euclidean_space set" assumes "AR S \ closedin (top_of_set U) S" shows "S retract_of U" using AR_imp_absolute_retract assms homeomorphic_refl by blast lemma AR_homeomorphic_AR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "AR T" "S homeomorphic T" shows "AR S" unfolding AR_def by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym) lemma homeomorphic_AR_iff_AR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" shows "S homeomorphic T \ AR S \ AR T" by (metis AR_homeomorphic_AR homeomorphic_sym) lemma ANR_imp_absolute_neighbourhood_extensor: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "ANR S" and contf: "continuous_on T f" and "f ` T \ S" and cloUT: "closedin (top_of_set U) T" obtains V g where "T \ V" "openin (top_of_set U) V" "continuous_on V g" "g ` V \ S" "\x. x \ T \ g x = f x" proof - have "aff_dim S < int (DIM('b \ real))" using aff_dim_le_DIM [of S] by simp then obtain C and S' :: "('b * real) set" where C: "convex C" "C \ {}" and cloCS: "closedin (top_of_set C) S'" and hom: "S homeomorphic S'" by (metis that homeomorphic_closedin_convex) then obtain D where opD: "openin (top_of_set C) D" and "S' retract_of D" using \ANR S\ by (auto simp: ANR_def) then obtain r where "S' \ D" and contr: "continuous_on D r" and "r ` D \ S'" and rid: "\x. x \ S' \ r x = x" by (auto simp: retraction_def retract_of_def) obtain g h where homgh: "homeomorphism S S' g h" using hom by (force simp: homeomorphic_def) have "continuous_on (f ` T) g" by (meson \f ` T \ S\ continuous_on_subset homeomorphism_def homgh) then have contgf: "continuous_on T (g \ f)" by (intro continuous_on_compose contf) have gfTC: "(g \ f) ` T \ C" proof - have "g ` S = S'" by (metis (no_types) homeomorphism_def homgh) then show ?thesis by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology) qed obtain f' where contf': "continuous_on U f'" and "f' ` U \ C" and eq: "\x. x \ T \ f' x = (g \ f) x" by (metis Dugundji [OF C cloUT contgf gfTC]) show ?thesis proof (rule_tac V = "U \ f' -` D" and g = "h \ r \ f'" in that) show "T \ U \ f' -` D" using cloUT closedin_imp_subset \S' \ D\ \f ` T \ S\ eq homeomorphism_image1 homgh by fastforce show ope: "openin (top_of_set U) (U \ f' -` D)" using \f' ` U \ C\ by (auto simp: opD contf' continuous_openin_preimage) have conth: "continuous_on (r ` f' ` (U \ f' -` D)) h" apply (rule continuous_on_subset [of S']) using homeomorphism_def homgh apply blast using \r ` D \ S'\ by blast show "continuous_on (U \ f' -` D) (h \ r \ f')" apply (intro continuous_on_compose conth continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto) done show "(h \ r \ f') ` (U \ f' -` D) \ S" using \homeomorphism S S' g h\ \f' ` U \ C\ \r ` D \ S'\ by (auto simp: homeomorphism_def) show "\x. x \ T \ (h \ r \ f') x = f x" using \homeomorphism S S' g h\ \f ` T \ S\ eq by (auto simp: rid homeomorphism_def) qed qed corollary ANR_imp_absolute_neighbourhood_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ANR S" "S homeomorphic S'" and clo: "closedin (top_of_set U) S'" obtains V where "openin (top_of_set U) V" "S' retract_of V" proof - obtain g h where hom: "homeomorphism S S' g h" using assms by (force simp: homeomorphic_def) have h: "continuous_on S' h" " h ` S' \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done from ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo] obtain V h' where "S' \ V" and opUV: "openin (top_of_set U) V" and h': "continuous_on V h'" "h' ` V \ S" and h'h:"\x. x \ S' \ h' x = h x" by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo]) have "S' retract_of V" proof (simp add: retraction_def retract_of_def, intro exI conjI \S' \ V\) show "continuous_on V (g \ h')" apply (intro continuous_on_compose h') apply (meson hom continuous_on_subset h' homeomorphism_cont1) done show "(g \ h') ` V \ S'" using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) show "\x\S'. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed then show ?thesis by (rule that [OF opUV]) qed corollary ANR_imp_absolute_neighbourhood_retract_UNIV: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'" obtains V where "open V" "S' retract_of V" using ANR_imp_absolute_neighbourhood_retract [OF \ANR S\ hom] by (metis clo closed_closedin open_openin subtopology_UNIV) corollary neighbourhood_extension_into_ANR: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and fim: "f ` S \ T" and "ANR T" "closed S" obtains V g where "S \ V" "open V" "continuous_on V g" "g ` V \ T" "\x. x \ S \ g x = f x" using ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf fim] by (metis \closed S\ closed_closedin open_openin subtopology_UNIV) lemma absolute_neighbourhood_extensor_imp_ANR: fixes S :: "'a::euclidean_space set" assumes "\f :: 'a * real \ 'a. \U T. \continuous_on T f; f ` T \ S; closedin (top_of_set U) T\ \ \V g. T \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)" shows "ANR S" proof (clarsimp simp: ANR_def) fix U and T :: "('a * real) set" assume "S homeomorphic T" and clo: "closedin (top_of_set U) T" then obtain g h where hom: "homeomorphism S T g h" by (force simp: homeomorphic_def) have h: "continuous_on T h" " h ` T \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done obtain V h' where "T \ V" and opV: "openin (top_of_set U) V" and h': "continuous_on V h'" "h' ` V \ S" and h'h: "\x\T. h' x = h x" using assms [OF h clo] by blast have [simp]: "T \ U" using clo closedin_imp_subset by auto have "T retract_of V" proof (simp add: retraction_def retract_of_def, intro exI conjI \T \ V\) show "continuous_on V (g \ h')" apply (intro continuous_on_compose h') apply (meson hom continuous_on_subset h' homeomorphism_cont1) done show "(g \ h') ` V \ T" using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) show "\x\T. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed then show "\V. openin (top_of_set U) V \ T retract_of V" using opV by blast qed lemma ANR_eq_absolute_neighbourhood_extensor: fixes S :: "'a::euclidean_space set" shows "ANR S \ (\f :: 'a * real \ 'a. \U T. continuous_on T f \ f ` T \ S \ closedin (top_of_set U) T \ (\V g. T \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)))" apply (rule iffI) apply (metis ANR_imp_absolute_neighbourhood_extensor) apply (simp add: absolute_neighbourhood_extensor_imp_ANR) done lemma ANR_imp_neighbourhood_retract: fixes S :: "'a::euclidean_space set" assumes "ANR S" "closedin (top_of_set U) S" obtains V where "openin (top_of_set U) V" "S retract_of V" using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast lemma ANR_imp_absolute_closed_neighbourhood_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ANR S" "S homeomorphic S'" and US': "closedin (top_of_set U) S'" obtains V W where "openin (top_of_set U) V" "closedin (top_of_set U) W" "S' \ V" "V \ W" "S' retract_of W" proof - obtain Z where "openin (top_of_set U) Z" and S'Z: "S' retract_of Z" by (blast intro: assms ANR_imp_absolute_neighbourhood_retract) then have UUZ: "closedin (top_of_set U) (U - Z)" by auto have "S' \ (U - Z) = {}" using \S' retract_of Z\ closedin_retract closedin_subtopology by fastforce then obtain V W where "openin (top_of_set U) V" and "openin (top_of_set U) W" and "S' \ V" "U - Z \ W" "V \ W = {}" using separation_normal_local [OF US' UUZ] by auto moreover have "S' retract_of U - W" apply (rule retract_of_subset [OF S'Z]) using US' \S' \ V\ \V \ W = {}\ closedin_subset apply fastforce using Diff_subset_conv \U - Z \ W\ by blast ultimately show ?thesis apply (rule_tac V=V and W = "U-W" in that) using openin_imp_subset apply force+ done qed lemma ANR_imp_closed_neighbourhood_retract: fixes S :: "'a::euclidean_space set" assumes "ANR S" "closedin (top_of_set U) S" obtains V W where "openin (top_of_set U) V" "closedin (top_of_set U) W" "S \ V" "V \ W" "S retract_of W" by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl) lemma ANR_homeomorphic_ANR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "ANR T" "S homeomorphic T" shows "ANR S" unfolding ANR_def by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym) lemma homeomorphic_ANR_iff_ANR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" shows "S homeomorphic T \ ANR S \ ANR T" by (metis ANR_homeomorphic_ANR homeomorphic_sym) subsection \Analogous properties of ENRs\ lemma ENR_imp_absolute_neighbourhood_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ENR S" and hom: "S homeomorphic S'" and "S' \ U" obtains V where "openin (top_of_set U) V" "S' retract_of V" proof - obtain X where "open X" "S retract_of X" using \ENR S\ by (auto simp: ENR_def) then obtain r where "retraction X S r" by (auto simp: retract_of_def) have "locally compact S'" using retract_of_locally_compact open_imp_locally_compact homeomorphic_local_compactness \S retract_of X\ \open X\ hom by blast then obtain W where UW: "openin (top_of_set U) W" and WS': "closedin (top_of_set W) S'" apply (rule locally_compact_closedin_open) apply (rename_tac W) apply (rule_tac W = "U \ W" in that, blast) by (simp add: \S' \ U\ closedin_limpt) obtain f g where hom: "homeomorphism S S' f g" using assms by (force simp: homeomorphic_def) have contg: "continuous_on S' g" using hom homeomorphism_def by blast moreover have "g ` S' \ S" by (metis hom equalityE homeomorphism_def) ultimately obtain h where conth: "continuous_on W h" and hg: "\x. x \ S' \ h x = g x" using Tietze_unbounded [of S' g W] WS' by blast have "W \ U" using UW openin_open by auto have "S' \ W" using WS' closedin_closed by auto have him: "\x. x \ S' \ h x \ X" by (metis (no_types) \S retract_of X\ hg hom homeomorphism_def image_insert insert_absorb insert_iff retract_of_imp_subset subset_eq) have "S' retract_of (W \ h -` X)" proof (simp add: retraction_def retract_of_def, intro exI conjI) show "S' \ W" "S' \ h -` X" using him WS' closedin_imp_subset by blast+ show "continuous_on (W \ h -` X) (f \ r \ h)" proof (intro continuous_on_compose) show "continuous_on (W \ h -` X) h" by (meson conth continuous_on_subset inf_le1) show "continuous_on (h ` (W \ h -` X)) r" proof - have "h ` (W \ h -` X) \ X" by blast then show "continuous_on (h ` (W \ h -` X)) r" by (meson \retraction X S r\ continuous_on_subset retraction) qed show "continuous_on (r ` h ` (W \ h -` X)) f" apply (rule continuous_on_subset [of S]) using hom homeomorphism_def apply blast apply clarify apply (meson \retraction X S r\ subsetD imageI retraction_def) done qed show "(f \ r \ h) ` (W \ h -` X) \ S'" using \retraction X S r\ hom by (auto simp: retraction_def homeomorphism_def) show "\x\S'. (f \ r \ h) x = x" using \retraction X S r\ hom by (auto simp: retraction_def homeomorphism_def hg) qed then show ?thesis apply (rule_tac V = "W \ h -` X" in that) apply (rule openin_trans [OF _ UW]) using \continuous_on W h\ \open X\ continuous_openin_preimage_eq apply blast+ done qed corollary ENR_imp_absolute_neighbourhood_retract_UNIV: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ENR S" "S homeomorphic S'" obtains T' where "open T'" "S' retract_of T'" by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV) lemma ENR_homeomorphic_ENR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "ENR T" "S homeomorphic T" shows "ENR S" unfolding ENR_def by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym) lemma homeomorphic_ENR_iff_ENR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" shows "ENR S \ ENR T" by (meson ENR_homeomorphic_ENR assms homeomorphic_sym) lemma ENR_translation: fixes S :: "'a::euclidean_space set" shows "ENR(image (\x. a + x) S) \ ENR S" by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR) lemma ENR_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "ENR (image f S) \ ENR S" apply (rule homeomorphic_ENR_iff_ENR) using assms homeomorphic_sym linear_homeomorphic_image by auto text \Some relations among the concepts. We also relate AR to being a retract of UNIV, which is often a more convenient proxy in the closed case.\ lemma AR_imp_ANR: "AR S \ ANR S" using ANR_def AR_def by fastforce lemma ENR_imp_ANR: fixes S :: "'a::euclidean_space set" shows "ENR S \ ANR S" apply (simp add: ANR_def) by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset) lemma ENR_ANR: fixes S :: "'a::euclidean_space set" shows "ENR S \ ANR S \ locally compact S" proof assume "ENR S" then have "locally compact S" using ENR_def open_imp_locally_compact retract_of_locally_compact by auto then show "ANR S \ locally compact S" using ENR_imp_ANR \ENR S\ by blast next assume "ANR S \ locally compact S" then have "ANR S" "locally compact S" by auto then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T" using locally_compact_homeomorphic_closed by (metis DIM_prod DIM_real Suc_eq_plus1 lessI) then show "ENR S" using \ANR S\ apply (simp add: ANR_def) apply (drule_tac x=UNIV in spec) apply (drule_tac x=T in spec, clarsimp) apply (meson ENR_def ENR_homeomorphic_ENR open_openin) done qed lemma AR_ANR: fixes S :: "'a::euclidean_space set" shows "AR S \ ANR S \ contractible S \ S \ {}" (is "?lhs = ?rhs") proof assume ?lhs obtain C and S' :: "('a * real) set" where "convex C" "C \ {}" "closedin (top_of_set C) S'" "S homeomorphic S'" apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"]) using aff_dim_le_DIM [of S] by auto with \AR S\ have "contractible S" apply (simp add: AR_def) apply (drule_tac x=C in spec) apply (drule_tac x="S'" in spec, simp) using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce with \AR S\ show ?rhs apply (auto simp: AR_imp_ANR) apply (force simp: AR_def) done next assume ?rhs then obtain a and h:: "real \ 'a \ 'a" where conth: "continuous_on ({0..1} \ S) h" and hS: "h ` ({0..1} \ S) \ S" and [simp]: "\x. h(0, x) = x" and [simp]: "\x. h(1, x) = a" and "ANR S" "S \ {}" by (auto simp: contractible_def homotopic_with_def) then have "a \ S" by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one) have "\g. continuous_on W g \ g ` W \ S \ (\x\T. g x = f x)" if f: "continuous_on T f" "f ` T \ S" and WT: "closedin (top_of_set W) T" for W T and f :: "'a \ real \ 'a" proof - obtain U g where "T \ U" and WU: "openin (top_of_set W) U" and contg: "continuous_on U g" and "g ` U \ S" and gf: "\x. x \ T \ g x = f x" using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \ANR S\, rule_format, OF f WT] by auto have WWU: "closedin (top_of_set W) (W - U)" using WU closedin_diff by fastforce moreover have "(W - U) \ T = {}" using \T \ U\ by auto ultimately obtain V V' where WV': "openin (top_of_set W) V'" and WV: "openin (top_of_set W) V" and "W - U \ V'" "T \ V" "V' \ V = {}" using separation_normal_local [of W "W-U" T] WT by blast then have WVT: "T \ (W - V) = {}" by auto have WWV: "closedin (top_of_set W) (W - V)" using WV closedin_diff by fastforce obtain j :: " 'a \ real \ real" where contj: "continuous_on W j" and j: "\x. x \ W \ j x \ {0..1}" and j0: "\x. x \ W - V \ j x = 1" and j1: "\x. x \ T \ j x = 0" by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment) have Weq: "W = (W - V) \ (W - V')" using \V' \ V = {}\ by force show ?thesis proof (intro conjI exI) have *: "continuous_on (W - V') (\x. h (j x, g x))" apply (rule continuous_on_compose2 [OF conth continuous_on_Pair]) apply (rule continuous_on_subset [OF contj Diff_subset]) apply (rule continuous_on_subset [OF contg]) apply (metis Diff_subset_conv Un_commute \W - U \ V'\) using j \g ` U \ S\ \W - U \ V'\ apply fastforce done show "continuous_on W (\x. if x \ W - V then a else h (j x, g x))" apply (subst Weq) apply (rule continuous_on_cases_local) apply (simp_all add: Weq [symmetric] WWV *) using WV' closedin_diff apply fastforce apply (auto simp: j0 j1) done next have "h (j (x, y), g (x, y)) \ S" if "(x, y) \ W" "(x, y) \ V" for x y proof - have "j(x, y) \ {0..1}" using j that by blast moreover have "g(x, y) \ S" using \V' \ V = {}\ \W - U \ V'\ \g ` U \ S\ that by fastforce ultimately show ?thesis using hS by blast qed with \a \ S\ \g ` U \ S\ show "(\x. if x \ W - V then a else h (j x, g x)) ` W \ S" by auto next show "\x\T. (if x \ W - V then a else h (j x, g x)) = f x" using \T \ V\ by (auto simp: j0 j1 gf) qed qed then show ?lhs by (simp add: AR_eq_absolute_extensor) qed lemma ANR_retract_of_ANR: fixes S :: "'a::euclidean_space set" assumes "ANR T" "S retract_of T" shows "ANR S" using assms apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def) apply (clarsimp elim!: all_forward) apply (erule impCE, metis subset_trans) apply (clarsimp elim!: ex_forward) apply (rule_tac x="r \ g" in exI) by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans) lemma AR_retract_of_AR: fixes S :: "'a::euclidean_space set" shows "\AR T; S retract_of T\ \ AR S" using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce lemma ENR_retract_of_ENR: "\ENR T; S retract_of T\ \ ENR S" by (meson ENR_def retract_of_trans) lemma retract_of_UNIV: fixes S :: "'a::euclidean_space set" shows "S retract_of UNIV \ AR S \ closed S" by (metis AR_ANR AR_imp_retract ENR_def ENR_imp_ANR closed_UNIV closed_closedin contractible_UNIV empty_not_UNIV open_UNIV retract_of_closed retract_of_contractible retract_of_empty(1) subtopology_UNIV) lemma compact_AR: fixes S :: "'a::euclidean_space set" shows "compact S \ AR S \ compact S \ S retract_of UNIV" using compact_imp_closed retract_of_UNIV by blast text \More properties of ARs, ANRs and ENRs\ lemma not_AR_empty [simp]: "\ AR({})" by (auto simp: AR_def) lemma ENR_empty [simp]: "ENR {}" by (simp add: ENR_def) lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)" by (simp add: ENR_imp_ANR) lemma convex_imp_AR: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ {}\ \ AR S" apply (rule absolute_extensor_imp_AR) apply (rule Dugundji, assumption+) by blast lemma convex_imp_ANR: fixes S :: "'a::euclidean_space set" shows "convex S \ ANR S" using ANR_empty AR_imp_ANR convex_imp_AR by blast lemma ENR_convex_closed: fixes S :: "'a::euclidean_space set" shows "\closed S; convex S\ \ ENR S" using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)" using retract_of_UNIV by auto lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)" by (simp add: AR_imp_ANR) lemma ENR_UNIV [simp]:"ENR UNIV" using ENR_def by blast lemma AR_singleton: fixes a :: "'a::euclidean_space" shows "AR {a}" using retract_of_UNIV by blast lemma ANR_singleton: fixes a :: "'a::euclidean_space" shows "ANR {a}" by (simp add: AR_imp_ANR AR_singleton) lemma ENR_singleton: "ENR {a}" using ENR_def by blast text \ARs closed under union\ lemma AR_closed_Un_local_aux: fixes U :: "'a::euclidean_space set" assumes "closedin (top_of_set U) S" "closedin (top_of_set U) T" "AR S" "AR T" "AR(S \ T)" shows "(S \ T) retract_of U" proof - have "S \ T \ {}" using assms AR_def by fastforce have "S \ U" "T \ U" using assms by (auto simp: closedin_imp_subset) define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" have US': "closedin (top_of_set U) S'" using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have UT': "closedin (top_of_set U) T'" using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have "S \ S'" using S'_def \S \ U\ setdist_sing_in_set by fastforce have "T \ T'" using T'_def \T \ U\ setdist_sing_in_set by fastforce have "S \ T \ W" "W \ U" using \S \ U\ by (auto simp: W_def setdist_sing_in_set) have "(S \ T) retract_of W" apply (rule AR_imp_absolute_retract [OF \AR(S \ T)\]) apply (simp add: homeomorphic_refl) apply (rule closedin_subset_trans [of U]) apply (simp_all add: assms closedin_Int \S \ T \ W\ \W \ U\) done then obtain r0 where "S \ T \ W" and contr0: "continuous_on W r0" and "r0 ` W \ S \ T" and r0 [simp]: "\x. x \ S \ T \ r0 x = x" by (auto simp: retract_of_def retraction_def) have ST: "x \ W \ x \ S \ x \ T" for x using setdist_eq_0_closedin \S \ T \ {}\ assms by (force simp: W_def setdist_sing_in_set) have "S' \ T' = W" by (auto simp: S'_def T'_def W_def) then have cloUW: "closedin (top_of_set U) W" using closedin_Int US' UT' by blast define r where "r \ \x. if x \ W then r0 x else x" have "r ` (W \ S) \ S" "r ` (W \ T) \ T" using \r0 ` W \ S \ T\ r_def by auto have contr: "continuous_on (W \ (S \ T)) r" unfolding r_def proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) show "closedin (top_of_set (W \ (S \ T))) W" using \S \ U\ \T \ U\ \W \ U\ \closedin (top_of_set U) W\ closedin_subset_trans by fastforce show "closedin (top_of_set (W \ (S \ T))) (S \ T)" by (meson \S \ U\ \T \ U\ \W \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) show "\x. x \ W \ x \ W \ x \ S \ T \ x \ W \ r0 x = x" by (auto simp: ST) qed have cloUWS: "closedin (top_of_set U) (W \ S)" by (simp add: cloUW assms closedin_Un) obtain g where contg: "continuous_on U g" and "g ` U \ S" and geqr: "\x. x \ W \ S \ g x = r x" apply (rule AR_imp_absolute_extensor [OF \AR S\ _ _ cloUWS]) apply (rule continuous_on_subset [OF contr]) using \r ` (W \ S) \ S\ apply auto done have cloUWT: "closedin (top_of_set U) (W \ T)" by (simp add: cloUW assms closedin_Un) obtain h where conth: "continuous_on U h" and "h ` U \ T" and heqr: "\x. x \ W \ T \ h x = r x" apply (rule AR_imp_absolute_extensor [OF \AR T\ _ _ cloUWT]) apply (rule continuous_on_subset [OF contr]) using \r ` (W \ T) \ T\ apply auto done have "U = S' \ T'" by (force simp: S'_def T'_def) then have cont: "continuous_on U (\x. if x \ S' then g x else h x)" apply (rule ssubst) apply (rule continuous_on_cases_local) using US' UT' \S' \ T' = W\ \U = S' \ T'\ contg conth continuous_on_subset geqr heqr apply auto done have UST: "(\x. if x \ S' then g x else h x) ` U \ S \ T" using \g ` U \ S\ \h ` U \ T\ by auto show ?thesis apply (simp add: retract_of_def retraction_def \S \ U\ \T \ U\) apply (rule_tac x="\x. if x \ S' then g x else h x" in exI) apply (intro conjI cont UST) by (metis IntI ST Un_iff \S \ S'\ \S' \ T' = W\ \T \ T'\ subsetD geqr heqr r0 r_def) qed lemma AR_closed_Un_local: fixes S :: "'a::euclidean_space set" assumes STS: "closedin (top_of_set (S \ T)) S" and STT: "closedin (top_of_set (S \ T)) T" and "AR S" "AR T" "AR(S \ T)" shows "AR(S \ T)" proof - have "C retract_of U" if hom: "S \ T homeomorphic C" and UC: "closedin (top_of_set U) C" for U and C :: "('a * real) set" proof - obtain f g where hom: "homeomorphism (S \ T) C f g" using hom by (force simp: homeomorphic_def) have US: "closedin (top_of_set U) (C \ g -` S)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) using hom homeomorphism_def apply blast apply (metis hom homeomorphism_def set_eq_subset) done have UT: "closedin (top_of_set U) (C \ g -` T)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) using hom homeomorphism_def apply blast apply (metis hom homeomorphism_def set_eq_subset) done have ARS: "AR (C \ g -` S)" apply (rule AR_homeomorphic_AR [OF \AR S\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have ART: "AR (C \ g -` T)" apply (rule AR_homeomorphic_AR [OF \AR T\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have ARI: "AR ((C \ g -` S) \ (C \ g -` T))" apply (rule AR_homeomorphic_AR [OF \AR (S \ T)\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have "C = (C \ g -` S) \ (C \ g -` T)" using hom by (auto simp: homeomorphism_def) then show ?thesis by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI]) qed then show ?thesis by (force simp: AR_def) qed corollary AR_closed_Un: fixes S :: "'a::euclidean_space set" shows "\closed S; closed T; AR S; AR T; AR (S \ T)\ \ AR (S \ T)" by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV) text \ANRs closed under union\ lemma ANR_closed_Un_local_aux: fixes U :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "ANR S" "ANR T" "ANR(S \ T)" obtains V where "openin (top_of_set U) V" "(S \ T) retract_of V" proof (cases "S = {} \ T = {}") case True with assms that show ?thesis by (metis ANR_imp_neighbourhood_retract Un_commute inf_bot_right sup_inf_absorb) next case False then have [simp]: "S \ {}" "T \ {}" by auto have "S \ U" "T \ U" using assms by (auto simp: closedin_imp_subset) define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" have cloUS': "closedin (top_of_set U) S'" using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have cloUT': "closedin (top_of_set U) T'" using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have "S \ S'" using S'_def \S \ U\ setdist_sing_in_set by fastforce have "T \ T'" using T'_def \T \ U\ setdist_sing_in_set by fastforce have "S' \ T' = U" by (auto simp: S'_def T'_def) have "W \ S'" by (simp add: Collect_mono S'_def W_def) have "W \ T'" by (simp add: Collect_mono T'_def W_def) have ST_W: "S \ T \ W" and "W \ U" using \S \ U\ by (force simp: W_def setdist_sing_in_set)+ have "S' \ T' = W" by (auto simp: S'_def T'_def W_def) then have cloUW: "closedin (top_of_set U) W" using closedin_Int cloUS' cloUT' by blast obtain W' W0 where "openin (top_of_set W) W'" and cloWW0: "closedin (top_of_set W) W0" and "S \ T \ W'" "W' \ W0" and ret: "(S \ T) retract_of W0" apply (rule ANR_imp_closed_neighbourhood_retract [OF \ANR(S \ T)\]) apply (rule closedin_subset_trans [of U, OF _ ST_W \W \ U\]) apply (blast intro: assms)+ done then obtain U0 where opeUU0: "openin (top_of_set U) U0" and U0: "S \ T \ U0" "U0 \ W \ W0" unfolding openin_open using \W \ U\ by blast have "W0 \ U" using \W \ U\ cloWW0 closedin_subset by fastforce obtain r0 where "S \ T \ W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \ S \ T" and r0 [simp]: "\x. x \ S \ T \ r0 x = x" using ret by (force simp: retract_of_def retraction_def) have ST: "x \ W \ x \ S \ x \ T" for x using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin) define r where "r \ \x. if x \ W0 then r0 x else x" have "r ` (W0 \ S) \ S" "r ` (W0 \ T) \ T" using \r0 ` W0 \ S \ T\ r_def by auto have contr: "continuous_on (W0 \ (S \ T)) r" unfolding r_def proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) show "closedin (top_of_set (W0 \ (S \ T))) W0" apply (rule closedin_subset_trans [of U]) using cloWW0 cloUW closedin_trans \W0 \ U\ \S \ U\ \T \ U\ apply blast+ done show "closedin (top_of_set (W0 \ (S \ T))) (S \ T)" by (meson \S \ U\ \T \ U\ \W0 \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) show "\x. x \ W0 \ x \ W0 \ x \ S \ T \ x \ W0 \ r0 x = x" using ST cloWW0 closedin_subset by fastforce qed have cloS'WS: "closedin (top_of_set S') (W0 \ S)" by (meson closedin_subset_trans US cloUS' \S \ S'\ \W \ S'\ cloUW cloWW0 closedin_Un closedin_imp_subset closedin_trans) obtain W1 g where "W0 \ S \ W1" and contg: "continuous_on W1 g" and opeSW1: "openin (top_of_set S') W1" and "g ` W1 \ S" and geqr: "\x. x \ W0 \ S \ g x = r x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ _ \r ` (W0 \ S) \ S\ cloS'WS]) apply (rule continuous_on_subset [OF contr], blast+) done have cloT'WT: "closedin (top_of_set T') (W0 \ T)" by (meson closedin_subset_trans UT cloUT' \T \ T'\ \W \ T'\ cloUW cloWW0 closedin_Un closedin_imp_subset closedin_trans) obtain W2 h where "W0 \ T \ W2" and conth: "continuous_on W2 h" and opeSW2: "openin (top_of_set T') W2" and "h ` W2 \ T" and heqr: "\x. x \ W0 \ T \ h x = r x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ _ \r ` (W0 \ T) \ T\ cloT'WT]) apply (rule continuous_on_subset [OF contr], blast+) done have "S' \ T' = W" by (force simp: S'_def T'_def W_def) obtain O1 O2 where "open O1" "W1 = S' \ O1" "open O2" "W2 = T' \ O2" using opeSW1 opeSW2 by (force simp: openin_open) show ?thesis proof have eq: "W1 - (W - U0) \ (W2 - (W - U0)) = ((U - T') \ O1 \ (U - S') \ O2 \ U \ O1 \ O2) - (W - U0)" using \U0 \ W \ W0\ \W0 \ S \ W1\ \W0 \ T \ W2\ by (auto simp: \S' \ T' = U\ [symmetric] \S' \ T' = W\ [symmetric] \W1 = S' \ O1\ \W2 = T' \ O2\) show "openin (top_of_set U) (W1 - (W - U0) \ (W2 - (W - U0)))" apply (subst eq) apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \open O1\ \open O2\, simp_all) done have cloW1: "closedin (top_of_set (W1 - (W - U0) \ (W2 - (W - U0)))) (W1 - (W - U0))" using cloUS' apply (simp add: closedin_closed) apply (erule ex_forward) using U0 \W0 \ S \ W1\ apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) done have cloW2: "closedin (top_of_set (W1 - (W - U0) \ (W2 - (W - U0)))) (W2 - (W - U0))" using cloUT' apply (simp add: closedin_closed) apply (erule ex_forward) using U0 \W0 \ T \ W2\ apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) done have *: "\x\S \ T. (if x \ S' then g x else h x) = x" using ST \S' \ T' = W\ cloT'WT closedin_subset geqr heqr apply (auto simp: r_def, fastforce) using \S \ S'\ \T \ T'\ \W0 \ S \ W1\ \W1 = S' \ O1\ by auto have "\r. continuous_on (W1 - (W - U0) \ (W2 - (W - U0))) r \ r ` (W1 - (W - U0) \ (W2 - (W - U0))) \ S \ T \ (\x\S \ T. r x = x)" apply (rule_tac x = "\x. if x \ S' then g x else h x" in exI) apply (intro conjI *) apply (rule continuous_on_cases_local [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]]) using \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = W\ \g ` W1 \ S\ \h ` W2 \ T\ apply auto using \U0 \ W \ W0\ \W0 \ S \ W1\ apply (fastforce simp add: geqr heqr)+ done then show "S \ T retract_of W1 - (W - U0) \ (W2 - (W - U0))" using \W0 \ S \ W1\ \W0 \ T \ W2\ ST opeUU0 U0 by (auto simp: retract_of_def retraction_def) qed qed lemma ANR_closed_Un_local: fixes S :: "'a::euclidean_space set" assumes STS: "closedin (top_of_set (S \ T)) S" and STT: "closedin (top_of_set (S \ T)) T" and "ANR S" "ANR T" "ANR(S \ T)" shows "ANR(S \ T)" proof - have "\T. openin (top_of_set U) T \ C retract_of T" if hom: "S \ T homeomorphic C" and UC: "closedin (top_of_set U) C" for U and C :: "('a * real) set" proof - obtain f g where hom: "homeomorphism (S \ T) C f g" using hom by (force simp: homeomorphic_def) have US: "closedin (top_of_set U) (C \ g -` S)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) using hom [unfolded homeomorphism_def] apply blast apply (metis hom homeomorphism_def set_eq_subset) done have UT: "closedin (top_of_set U) (C \ g -` T)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) using hom [unfolded homeomorphism_def] apply blast apply (metis hom homeomorphism_def set_eq_subset) done have ANRS: "ANR (C \ g -` S)" apply (rule ANR_homeomorphic_ANR [OF \ANR S\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have ANRT: "ANR (C \ g -` T)" apply (rule ANR_homeomorphic_ANR [OF \ANR T\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have ANRI: "ANR ((C \ g -` S) \ (C \ g -` T))" apply (rule ANR_homeomorphic_ANR [OF \ANR (S \ T)\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have "C = (C \ g -` S) \ (C \ g -` T)" using hom by (auto simp: homeomorphism_def) then show ?thesis by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI]) qed then show ?thesis by (auto simp: ANR_def) qed corollary ANR_closed_Un: fixes S :: "'a::euclidean_space set" shows "\closed S; closed T; ANR S; ANR T; ANR (S \ T)\ \ ANR (S \ T)" by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int) lemma ANR_openin: fixes S :: "'a::euclidean_space set" assumes "ANR T" and opeTS: "openin (top_of_set T) S" shows "ANR S" proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) fix f :: "'a \ real \ 'a" and U C assume contf: "continuous_on C f" and fim: "f ` C \ S" and cloUC: "closedin (top_of_set U) C" have "f ` C \ T" using fim opeTS openin_imp_subset by blast obtain W g where "C \ W" and UW: "openin (top_of_set U) W" and contg: "continuous_on W g" and gim: "g ` W \ T" and geq: "\x. x \ C \ g x = f x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf \f ` C \ T\ cloUC]) using fim by auto show "\V g. C \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x\C. g x = f x)" proof (intro exI conjI) show "C \ W \ g -` S" using \C \ W\ fim geq by blast show "openin (top_of_set U) (W \ g -` S)" by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans) show "continuous_on (W \ g -` S) g" by (blast intro: continuous_on_subset [OF contg]) show "g ` (W \ g -` S) \ S" using gim by blast show "\x\C. g x = f x" using geq by blast qed qed lemma ENR_openin: fixes S :: "'a::euclidean_space set" assumes "ENR T" and opeTS: "openin (top_of_set T) S" shows "ENR S" using assms apply (simp add: ENR_ANR) using ANR_openin locally_open_subset by blast lemma ANR_neighborhood_retract: fixes S :: "'a::euclidean_space set" assumes "ANR U" "S retract_of T" "openin (top_of_set U) T" shows "ANR S" using ANR_openin ANR_retract_of_ANR assms by blast lemma ENR_neighborhood_retract: fixes S :: "'a::euclidean_space set" assumes "ENR U" "S retract_of T" "openin (top_of_set U) T" shows "ENR S" using ENR_openin ENR_retract_of_ENR assms by blast lemma ANR_rel_interior: fixes S :: "'a::euclidean_space set" shows "ANR S \ ANR(rel_interior S)" by (blast intro: ANR_openin openin_set_rel_interior) lemma ANR_delete: fixes S :: "'a::euclidean_space set" shows "ANR S \ ANR(S - {a})" by (blast intro: ANR_openin openin_delete openin_subtopology_self) lemma ENR_rel_interior: fixes S :: "'a::euclidean_space set" shows "ENR S \ ENR(rel_interior S)" by (blast intro: ENR_openin openin_set_rel_interior) lemma ENR_delete: fixes S :: "'a::euclidean_space set" shows "ENR S \ ENR(S - {a})" by (blast intro: ENR_openin openin_delete openin_subtopology_self) lemma open_imp_ENR: "open S \ ENR S" using ENR_def by blast lemma open_imp_ANR: fixes S :: "'a::euclidean_space set" shows "open S \ ANR S" by (simp add: ENR_imp_ANR open_imp_ENR) lemma ANR_ball [iff]: fixes a :: "'a::euclidean_space" shows "ANR(ball a r)" by (simp add: convex_imp_ANR) lemma ENR_ball [iff]: "ENR(ball a r)" by (simp add: open_imp_ENR) lemma AR_ball [simp]: fixes a :: "'a::euclidean_space" shows "AR(ball a r) \ 0 < r" by (auto simp: AR_ANR convex_imp_contractible) lemma ANR_cball [iff]: fixes a :: "'a::euclidean_space" shows "ANR(cball a r)" by (simp add: convex_imp_ANR) lemma ENR_cball: fixes a :: "'a::euclidean_space" shows "ENR(cball a r)" using ENR_convex_closed by blast lemma AR_cball [simp]: fixes a :: "'a::euclidean_space" shows "AR(cball a r) \ 0 \ r" by (auto simp: AR_ANR convex_imp_contractible) lemma ANR_box [iff]: fixes a :: "'a::euclidean_space" shows "ANR(cbox a b)" "ANR(box a b)" by (auto simp: convex_imp_ANR open_imp_ANR) lemma ENR_box [iff]: fixes a :: "'a::euclidean_space" shows "ENR(cbox a b)" "ENR(box a b)" apply (simp add: ENR_convex_closed closed_cbox) by (simp add: open_box open_imp_ENR) lemma AR_box [simp]: "AR(cbox a b) \ cbox a b \ {}" "AR(box a b) \ box a b \ {}" by (auto simp: AR_ANR convex_imp_contractible) lemma ANR_interior: fixes S :: "'a::euclidean_space set" shows "ANR(interior S)" by (simp add: open_imp_ANR) lemma ENR_interior: fixes S :: "'a::euclidean_space set" shows "ENR(interior S)" by (simp add: open_imp_ENR) lemma AR_imp_contractible: fixes S :: "'a::euclidean_space set" shows "AR S \ contractible S" by (simp add: AR_ANR) lemma ENR_imp_locally_compact: fixes S :: "'a::euclidean_space set" shows "ENR S \ locally compact S" by (simp add: ENR_ANR) lemma ANR_imp_locally_path_connected: fixes S :: "'a::euclidean_space set" assumes "ANR S" shows "locally path_connected S" proof - obtain U and T :: "('a \ real) set" where "convex U" "U \ {}" and UT: "closedin (top_of_set U) T" and "S homeomorphic T" apply (rule homeomorphic_closedin_convex [of S]) using aff_dim_le_DIM [of S] apply auto done then have "locally path_connected T" by (meson ANR_imp_absolute_neighbourhood_retract assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected) then have S: "locally path_connected S" if "openin (top_of_set U) V" "T retract_of V" "U \ {}" for V using \S homeomorphic T\ homeomorphic_locally homeomorphic_path_connectedness by blast show ?thesis using assms apply (clarsimp simp: ANR_def) apply (drule_tac x=U in spec) apply (drule_tac x=T in spec) using \S homeomorphic T\ \U \ {}\ UT apply (blast intro: S) done qed lemma ANR_imp_locally_connected: fixes S :: "'a::euclidean_space set" assumes "ANR S" shows "locally connected S" using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto lemma AR_imp_locally_path_connected: fixes S :: "'a::euclidean_space set" assumes "AR S" shows "locally path_connected S" by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms) lemma AR_imp_locally_connected: fixes S :: "'a::euclidean_space set" assumes "AR S" shows "locally connected S" using ANR_imp_locally_connected AR_ANR assms by blast lemma ENR_imp_locally_path_connected: fixes S :: "'a::euclidean_space set" assumes "ENR S" shows "locally path_connected S" by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms) lemma ENR_imp_locally_connected: fixes S :: "'a::euclidean_space set" assumes "ENR S" shows "locally connected S" using ANR_imp_locally_connected ENR_ANR assms by blast lemma ANR_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "ANR S" "ANR T" shows "ANR(S \ T)" proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) fix f :: " ('a \ 'b) \ real \ 'a \ 'b" and U C assume "continuous_on C f" and fim: "f ` C \ S \ T" and cloUC: "closedin (top_of_set U) C" have contf1: "continuous_on C (fst \ f)" by (simp add: \continuous_on C f\ continuous_on_fst) obtain W1 g where "C \ W1" and UW1: "openin (top_of_set U) W1" and contg: "continuous_on W1 g" and gim: "g ` W1 \ S" and geq: "\x. x \ C \ g x = (fst \ f) x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ contf1 _ cloUC]) using fim apply auto done have contf2: "continuous_on C (snd \ f)" by (simp add: \continuous_on C f\ continuous_on_snd) obtain W2 h where "C \ W2" and UW2: "openin (top_of_set U) W2" and conth: "continuous_on W2 h" and him: "h ` W2 \ T" and heq: "\x. x \ C \ h x = (snd \ f) x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf2 _ cloUC]) using fim apply auto done show "\V g. C \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ T \ (\x\C. g x = f x)" proof (intro exI conjI) show "C \ W1 \ W2" by (simp add: \C \ W1\ \C \ W2\) show "openin (top_of_set U) (W1 \ W2)" by (simp add: UW1 UW2 openin_Int) show "continuous_on (W1 \ W2) (\x. (g x, h x))" by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1) show "(\x. (g x, h x)) ` (W1 \ W2) \ S \ T" using gim him by blast show "(\x\C. (g x, h x) = f x)" using geq heq by auto qed qed lemma AR_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "AR S" "AR T" shows "AR(S \ T)" using assms by (simp add: AR_ANR ANR_Times contractible_Times) (* Unused subsection\<^marker>\tag unimportant\\Retracts and intervals in ordered euclidean space\ lemma ANR_interval [iff]: fixes a :: "'a::ordered_euclidean_space" shows "ANR{a..b}" by (simp add: interval_cbox) lemma ENR_interval [iff]: fixes a :: "'a::ordered_euclidean_space" shows "ENR{a..b}" by (auto simp: interval_cbox) *) subsection \More advanced properties of ANRs and ENRs\ lemma ENR_rel_frontier_convex: fixes S :: "'a::euclidean_space set" assumes "bounded S" "convex S" shows "ENR(rel_frontier S)" proof (cases "S = {}") case True then show ?thesis by simp next case False with assms have "rel_interior S \ {}" by (simp add: rel_interior_eq_empty) then obtain a where a: "a \ rel_interior S" by auto have ahS: "affine hull S - {a} \ {x. closest_point (affine hull S) x \ a}" by (auto simp: closest_point_self) have "rel_frontier S retract_of affine hull S - {a}" by (simp add: assms a rel_frontier_retract_of_punctured_affine_hull) also have "\ retract_of {x. closest_point (affine hull S) x \ a}" apply (simp add: retract_of_def retraction_def ahS) apply (rule_tac x="closest_point (affine hull S)" in exI) apply (auto simp: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point) done finally have "rel_frontier S retract_of {x. closest_point (affine hull S) x \ a}" . moreover have "openin (top_of_set UNIV) (UNIV \ closest_point (affine hull S) -` (- {a}))" apply (rule continuous_openin_preimage_gen) apply (auto simp: False affine_imp_convex continuous_on_closest_point) done ultimately show ?thesis unfolding ENR_def apply (rule_tac x = "closest_point (affine hull S) -` (- {a})" in exI) apply (simp add: vimage_def) done qed lemma ANR_rel_frontier_convex: fixes S :: "'a::euclidean_space set" assumes "bounded S" "convex S" shows "ANR(rel_frontier S)" by (simp add: ENR_imp_ANR ENR_rel_frontier_convex assms) lemma ENR_closedin_Un_local: fixes S :: "'a::euclidean_space set" shows "\ENR S; ENR T; ENR(S \ T); closedin (top_of_set (S \ T)) S; closedin (top_of_set (S \ T)) T\ \ ENR(S \ T)" by (simp add: ENR_ANR ANR_closed_Un_local locally_compact_closedin_Un) lemma ENR_closed_Un: fixes S :: "'a::euclidean_space set" shows "\closed S; closed T; ENR S; ENR T; ENR(S \ T)\ \ ENR(S \ T)" by (auto simp: closed_subset ENR_closedin_Un_local) lemma absolute_retract_Un: fixes S :: "'a::euclidean_space set" shows "\S retract_of UNIV; T retract_of UNIV; (S \ T) retract_of UNIV\ \ (S \ T) retract_of UNIV" by (meson AR_closed_Un_local_aux closed_subset retract_of_UNIV retract_of_imp_subset) lemma retract_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" and Un: "(S \ T) retract_of U" and Int: "(S \ T) retract_of T" shows "S retract_of U" proof - obtain r where r: "continuous_on T r" "r ` T \ S \ T" "\x\S \ T. r x = x" using Int by (auto simp: retraction_def retract_of_def) have "S retract_of S \ T" unfolding retraction_def retract_of_def proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then x else r x)" apply (rule continuous_on_cases_local [OF clS clT]) using r by (auto) qed (use r in auto) also have "\ retract_of U" by (rule Un) finally show ?thesis . qed lemma AR_from_Un_Int_local: fixes S :: "'a::euclidean_space set" assumes clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" and Un: "AR(S \ T)" and Int: "AR(S \ T)" shows "AR S" apply (rule AR_retract_of_AR [OF Un]) by (meson AR_imp_retract clS clT closedin_closed_subset local.Int retract_from_Un_Int retract_of_refl sup_ge2) lemma AR_from_Un_Int_local': fixes S :: "'a::euclidean_space set" assumes "closedin (top_of_set (S \ T)) S" and "closedin (top_of_set (S \ T)) T" and "AR(S \ T)" "AR(S \ T)" shows "AR T" using AR_from_Un_Int_local [of T S] assms by (simp add: Un_commute Int_commute) lemma AR_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes clo: "closed S" "closed T" and Un: "AR(S \ T)" and Int: "AR(S \ T)" shows "AR S" by (metis AR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) lemma ANR_from_Un_Int_local: fixes S :: "'a::euclidean_space set" assumes clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" and Un: "ANR(S \ T)" and Int: "ANR(S \ T)" shows "ANR S" proof - obtain V where clo: "closedin (top_of_set (S \ T)) (S \ T)" and ope: "openin (top_of_set (S \ T)) V" and ret: "S \ T retract_of V" using ANR_imp_neighbourhood_retract [OF Int] by (metis clS clT closedin_Int) then obtain r where r: "continuous_on V r" and rim: "r ` V \ S \ T" and req: "\x\S \ T. r x = x" by (auto simp: retraction_def retract_of_def) have Vsub: "V \ S \ T" by (meson ope openin_contains_cball) have Vsup: "S \ T \ V" by (simp add: retract_of_imp_subset ret) then have eq: "S \ V = ((S \ T) - T) \ V" by auto have eq': "S \ V = S \ (V \ T)" using Vsub by blast have "continuous_on (S \ V \ T) (\x. if x \ S then x else r x)" proof (rule continuous_on_cases_local) show "closedin (top_of_set (S \ V \ T)) S" using clS closedin_subset_trans inf.boundedE by blast show "closedin (top_of_set (S \ V \ T)) (V \ T)" using clT Vsup by (auto simp: closedin_closed) show "continuous_on (V \ T) r" by (meson Int_lower1 continuous_on_subset r) qed (use req continuous_on_id in auto) with rim have "S retract_of S \ V" unfolding retraction_def retract_of_def apply (rule_tac x="\x. if x \ S then x else r x" in exI) apply (auto simp: eq') done then show ?thesis using ANR_neighborhood_retract [OF Un] using \S \ V = S \ T - T \ V\ clT ope by fastforce qed lemma ANR_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes clo: "closed S" "closed T" and Un: "ANR(S \ T)" and Int: "ANR(S \ T)" shows "ANR S" by (metis ANR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) lemma ANR_finite_Union_convex_closed: fixes \ :: "'a::euclidean_space set set" assumes \: "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" shows "ANR(\\)" proof - have "ANR(\\)" if "card \ < n" for n using assms that proof (induction n arbitrary: \) case 0 then show ?case by simp next case (Suc n) have "ANR(\\)" if "finite \" "\ \ \" for \ using that proof (induction \) case empty then show ?case by simp next case (insert C \) have "ANR (C \ \\)" proof (rule ANR_closed_Un) show "ANR (C \ \\)" unfolding Int_Union proof (rule Suc) show "finite ((\) C ` \)" by (simp add: insert.hyps(1)) show "\Ca. Ca \ (\) C ` \ \ closed Ca" by (metis (no_types, hide_lams) Suc.prems(2) closed_Int subsetD imageE insert.prems insertI1 insertI2) show "\Ca. Ca \ (\) C ` \ \ convex Ca" by (metis (mono_tags, lifting) Suc.prems(3) convex_Int imageE insert.prems insert_subset subsetCE) show "card ((\) C ` \) < n" proof - have "card \ \ n" by (meson Suc.prems(4) not_less not_less_eq) then show ?thesis by (metis Suc.prems(1) card_image_le card_seteq insert.hyps insert.prems insert_subset le_trans not_less) qed qed show "closed (\\)" using Suc.prems(2) insert.hyps(1) insert.prems by blast qed (use Suc.prems convex_imp_ANR insert.prems insert.IH in auto) then show ?case by simp qed then show ?case using Suc.prems(1) by blast qed then show ?thesis by blast qed lemma finite_imp_ANR: fixes S :: "'a::euclidean_space set" assumes "finite S" shows "ANR S" proof - have "ANR(\x \ S. {x})" by (blast intro: ANR_finite_Union_convex_closed assms) then show ?thesis by simp qed lemma ANR_insert: fixes S :: "'a::euclidean_space set" assumes "ANR S" "closed S" shows "ANR(insert a S)" by (metis ANR_closed_Un ANR_empty ANR_singleton Diff_disjoint Diff_insert_absorb assms closed_singleton insert_absorb insert_is_Un) lemma ANR_path_component_ANR: fixes S :: "'a::euclidean_space set" shows "ANR S \ ANR(path_component_set S x)" using ANR_imp_locally_path_connected ANR_openin openin_path_component_locally_path_connected by blast lemma ANR_connected_component_ANR: fixes S :: "'a::euclidean_space set" shows "ANR S \ ANR(connected_component_set S x)" by (metis ANR_openin openin_connected_component_locally_connected ANR_imp_locally_connected) lemma ANR_component_ANR: fixes S :: "'a::euclidean_space set" assumes "ANR S" "c \ components S" shows "ANR c" by (metis ANR_connected_component_ANR assms componentsE) subsection\Original ANR material, now for ENRs\ lemma ENR_bounded: fixes S :: "'a::euclidean_space set" assumes "bounded S" shows "ENR S \ (\U. open U \ bounded U \ S retract_of U)" (is "?lhs = ?rhs") proof obtain r where "0 < r" and r: "S \ ball 0 r" using bounded_subset_ballD assms by blast assume ?lhs then show ?rhs apply (clarsimp simp: ENR_def) apply (rule_tac x="ball 0 r \ U" in exI, auto) using r retract_of_imp_subset retract_of_subset by fastforce next assume ?rhs then show ?lhs using ENR_def by blast qed lemma absolute_retract_imp_AR_gen: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "S retract_of T" "convex T" "T \ {}" "S homeomorphic S'" "closedin (top_of_set U) S'" shows "S' retract_of U" proof - have "AR T" by (simp add: assms convex_imp_AR) then have "AR S" using AR_retract_of_AR assms by auto then show ?thesis using assms AR_imp_absolute_retract by metis qed lemma absolute_retract_imp_AR: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "S retract_of UNIV" "S homeomorphic S'" "closed S'" shows "S' retract_of UNIV" using AR_imp_absolute_retract_UNIV assms retract_of_UNIV by blast lemma homeomorphic_compact_arness: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "S homeomorphic S'" shows "compact S \ S retract_of UNIV \ compact S' \ S' retract_of UNIV" using assms homeomorphic_compactness apply auto apply (meson assms compact_AR homeomorphic_AR_iff_AR homeomorphic_compactness)+ done lemma absolute_retract_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes "(S \ T) retract_of UNIV" "(S \ T) retract_of UNIV" "closed S" "closed T" shows "S retract_of UNIV" using AR_from_Un_Int assms retract_of_UNIV by auto lemma ENR_from_Un_Int_gen: fixes S :: "'a::euclidean_space set" assumes "closedin (top_of_set (S \ T)) S" "closedin (top_of_set (S \ T)) T" "ENR(S \ T)" "ENR(S \ T)" shows "ENR S" apply (simp add: ENR_ANR) using ANR_from_Un_Int_local ENR_ANR assms locally_compact_closedin by blast lemma ENR_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "ENR(S \ T)" "ENR(S \ T)" shows "ENR S" by (meson ENR_from_Un_Int_gen assms closed_subset sup_ge1 sup_ge2) lemma ENR_finite_Union_convex_closed: fixes \ :: "'a::euclidean_space set set" assumes \: "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" shows "ENR(\ \)" by (simp add: ENR_ANR ANR_finite_Union_convex_closed \ clo closed_Union closed_imp_locally_compact con) lemma finite_imp_ENR: fixes S :: "'a::euclidean_space set" shows "finite S \ ENR S" by (simp add: ENR_ANR finite_imp_ANR finite_imp_closed closed_imp_locally_compact) lemma ENR_insert: fixes S :: "'a::euclidean_space set" assumes "closed S" "ENR S" shows "ENR(insert a S)" proof - have "ENR ({a} \ S)" by (metis ANR_insert ENR_ANR Un_commute Un_insert_right assms closed_imp_locally_compact closed_insert sup_bot_right) then show ?thesis by auto qed lemma ENR_path_component_ENR: fixes S :: "'a::euclidean_space set" assumes "ENR S" shows "ENR(path_component_set S x)" by (metis ANR_imp_locally_path_connected ENR_empty ENR_imp_ANR ENR_openin assms locally_path_connected_2 openin_subtopology_self path_component_eq_empty) (*UNUSED lemma ENR_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "ENR S" "ENR T" shows "ENR(S \ T)" using assms apply (simp add: ENR_ANR ANR_Times) thm locally_compact_Times oops SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);; *) subsection\Finally, spheres are ANRs and ENRs\ lemma absolute_retract_homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and U :: "'b::euclidean_space set" assumes "S homeomorphic U" "S \ {}" "S \ T" "convex U" "compact U" shows "S retract_of T" by (metis UNIV_I assms compact_AR convex_imp_AR homeomorphic_AR_iff_AR homeomorphic_compactness homeomorphic_empty(1) retract_of_subset subsetI) lemma frontier_retract_of_punctured_universe: fixes S :: "'a::euclidean_space set" assumes "convex S" "bounded S" "a \ interior S" shows "(frontier S) retract_of (- {a})" using rel_frontier_retract_of_punctured_affine_hull by (metis Compl_eq_Diff_UNIV affine_hull_nonempty_interior assms empty_iff rel_frontier_frontier rel_interior_nonempty_interior) lemma sphere_retract_of_punctured_universe_gen: fixes a :: "'a::euclidean_space" assumes "b \ ball a r" shows "sphere a r retract_of (- {b})" proof - have "frontier (cball a r) retract_of (- {b})" apply (rule frontier_retract_of_punctured_universe) using assms by auto then show ?thesis by simp qed lemma sphere_retract_of_punctured_universe: fixes a :: "'a::euclidean_space" assumes "0 < r" shows "sphere a r retract_of (- {a})" by (simp add: assms sphere_retract_of_punctured_universe_gen) lemma ENR_sphere: fixes a :: "'a::euclidean_space" shows "ENR(sphere a r)" proof (cases "0 < r") case True then have "sphere a r retract_of -{a}" by (simp add: sphere_retract_of_punctured_universe) with open_delete show ?thesis by (auto simp: ENR_def) next case False then show ?thesis using finite_imp_ENR by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial) qed corollary\<^marker>\tag unimportant\ ANR_sphere: fixes a :: "'a::euclidean_space" shows "ANR(sphere a r)" by (simp add: ENR_imp_ANR ENR_sphere) subsection\Spheres are connected, etc\ lemma locally_path_connected_sphere_gen: fixes S :: "'a::euclidean_space set" assumes "bounded S" and "convex S" shows "locally path_connected (rel_frontier S)" proof (cases "rel_interior S = {}") case True with assms show ?thesis by (simp add: rel_interior_eq_empty) next case False then obtain a where a: "a \ rel_interior S" by blast show ?thesis proof (rule retract_of_locally_path_connected) show "locally path_connected (affine hull S - {a})" by (meson convex_affine_hull convex_imp_locally_path_connected locally_open_subset openin_delete openin_subtopology_self) show "rel_frontier S retract_of affine hull S - {a}" using a assms rel_frontier_retract_of_punctured_affine_hull by blast qed qed lemma locally_connected_sphere_gen: fixes S :: "'a::euclidean_space set" assumes "bounded S" and "convex S" shows "locally connected (rel_frontier S)" by (simp add: ANR_imp_locally_connected ANR_rel_frontier_convex assms) lemma locally_path_connected_sphere: fixes a :: "'a::euclidean_space" shows "locally path_connected (sphere a r)" using ENR_imp_locally_path_connected ENR_sphere by blast lemma locally_connected_sphere: fixes a :: "'a::euclidean_space" shows "locally connected(sphere a r)" using ANR_imp_locally_connected ANR_sphere by blast subsection\Borsuk homotopy extension theorem\ text\It's only this late so we can use the concept of retraction, saying that the domain sets or range set are ENRs.\ theorem Borsuk_homotopy_extension_homotopic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes cloTS: "closedin (top_of_set T) S" and anr: "(ANR S \ ANR T) \ ANR U" and contf: "continuous_on T f" and "f ` T \ U" and "homotopic_with_canon (\x. True) S U f g" obtains g' where "homotopic_with_canon (\x. True) T U f g'" "continuous_on T g'" "image g' T \ U" "\x. x \ S \ g' x = g x" proof - have "S \ T" using assms closedin_imp_subset by blast obtain h where conth: "continuous_on ({0..1} \ S) h" and him: "h ` ({0..1} \ S) \ U" and [simp]: "\x. h(0, x) = f x" "\x. h(1::real, x) = g x" using assms by (auto simp: homotopic_with_def) define h' where "h' \ \z. if snd z \ S then h z else (f \ snd) z" define B where "B \ {0::real} \ T \ {0..1} \ S" have clo0T: "closedin (top_of_set ({0..1} \ T)) ({0::real} \ T)" by (simp add: Abstract_Topology.closedin_Times) moreover have cloT1S: "closedin (top_of_set ({0..1} \ T)) ({0..1} \ S)" by (simp add: Abstract_Topology.closedin_Times assms) ultimately have clo0TB:"closedin (top_of_set ({0..1} \ T)) B" by (auto simp: B_def) have cloBS: "closedin (top_of_set B) ({0..1} \ S)" by (metis (no_types) Un_subset_iff B_def closedin_subset_trans [OF cloT1S] clo0TB closedin_imp_subset closedin_self) moreover have cloBT: "closedin (top_of_set B) ({0} \ T)" using \S \ T\ closedin_subset_trans [OF clo0T] by (metis B_def Un_upper1 clo0TB closedin_closed inf_le1) moreover have "continuous_on ({0} \ T) (f \ snd)" apply (rule continuous_intros)+ apply (simp add: contf) done ultimately have conth': "continuous_on B h'" apply (simp add: h'_def B_def Un_commute [of "{0} \ T"]) apply (auto intro!: continuous_on_cases_local conth) done have "image h' B \ U" using \f ` T \ U\ him by (auto simp: h'_def B_def) obtain V k where "B \ V" and opeTV: "openin (top_of_set ({0..1} \ T)) V" and contk: "continuous_on V k" and kim: "k ` V \ U" and keq: "\x. x \ B \ k x = h' x" using anr proof assume ST: "ANR S \ ANR T" have eq: "({0} \ T \ {0..1} \ S) = {0::real} \ S" using \S \ T\ by auto have "ANR B" apply (simp add: B_def) apply (rule ANR_closed_Un_local) apply (metis cloBT B_def) apply (metis Un_commute cloBS B_def) apply (simp_all add: ANR_Times convex_imp_ANR ANR_singleton ST eq) done note Vk = that have *: thesis if "openin (top_of_set ({0..1::real} \ T)) V" "retraction V B r" for V r using that apply (clarsimp simp add: retraction_def) apply (rule Vk [of V "h' \ r"], assumption+) apply (metis continuous_on_compose conth' continuous_on_subset) using \h' ` B \ U\ apply force+ done show thesis apply (rule ANR_imp_neighbourhood_retract [OF \ANR B\ clo0TB]) apply (auto simp: ANR_Times ANR_singleton ST retract_of_def *) done next assume "ANR U" with ANR_imp_absolute_neighbourhood_extensor \h' ` B \ U\ clo0TB conth' that show ?thesis by blast qed define S' where "S' \ {x. \u::real. u \ {0..1} \ (u, x::'a) \ {0..1} \ T - V}" have "closedin (top_of_set T) S'" unfolding S'_def apply (rule closedin_compact_projection, blast) using closedin_self opeTV by blast have S'_def: "S' = {x. \u::real. (u, x::'a) \ {0..1} \ T - V}" by (auto simp: S'_def) have cloTS': "closedin (top_of_set T) S'" using S'_def \closedin (top_of_set T) S'\ by blast have "S \ S' = {}" using S'_def B_def \B \ V\ by force obtain a :: "'a \ real" where conta: "continuous_on T a" and "\x. x \ T \ a x \ closed_segment 1 0" and a1: "\x. x \ S \ a x = 1" and a0: "\x. x \ S' \ a x = 0" apply (rule Urysohn_local [OF cloTS cloTS' \S \ S' = {}\, of 1 0], blast) done then have ain: "\x. x \ T \ a x \ {0..1}" using closed_segment_eq_real_ivl by auto have inV: "(u * a t, t) \ V" if "t \ T" "0 \ u" "u \ 1" for t u proof (rule ccontr) assume "(u * a t, t) \ V" with ain [OF \t \ T\] have "a t = 0" apply simp apply (rule a0) by (metis (no_types, lifting) Diff_iff S'_def SigmaI atLeastAtMost_iff mem_Collect_eq mult_le_one mult_nonneg_nonneg that) show False using B_def \(u * a t, t) \ V\ \B \ V\ \a t = 0\ that by auto qed show ?thesis proof show hom: "homotopic_with_canon (\x. True) T U f (\x. k (a x, x))" proof (simp add: homotopic_with, intro exI conjI) show "continuous_on ({0..1} \ T) (k \ (\z. (fst z *\<^sub>R (a \ snd) z, snd z)))" apply (intro continuous_on_compose continuous_intros) apply (rule continuous_on_subset [OF conta], force) apply (rule continuous_on_subset [OF contk]) apply (force intro: inV) done show "(k \ (\z. (fst z *\<^sub>R (a \ snd) z, snd z))) ` ({0..1} \ T) \ U" using inV kim by auto show "\x\T. (k \ (\z. (fst z *\<^sub>R (a \ snd) z, snd z))) (0, x) = f x" by (simp add: B_def h'_def keq) show "\x\T. (k \ (\z. (fst z *\<^sub>R (a \ snd) z, snd z))) (1, x) = k (a x, x)" by auto qed show "continuous_on T (\x. k (a x, x))" using homotopic_with_imp_continuous_maps [OF hom] by auto show "(\x. k (a x, x)) ` T \ U" proof clarify fix t assume "t \ T" show "k (a t, t) \ U" by (metis \t \ T\ image_subset_iff inV kim not_one_le_zero linear mult_cancel_right1) qed show "\x. x \ S \ k (a x, x) = g x" by (simp add: B_def a1 h'_def keq) qed qed corollary\<^marker>\tag unimportant\ nullhomotopic_into_ANR_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" and "ANR T" and fim: "f ` S \ T" and "S \ {}" shows "(\c. homotopic_with_canon (\x. True) S T f (\x. c)) \ (\g. continuous_on UNIV g \ range g \ T \ (\x \ S. g x = f x))" (is "?lhs = ?rhs") proof assume ?lhs then obtain c where c: "homotopic_with_canon (\x. True) S T (\x. c) f" by (blast intro: homotopic_with_symD) have "closedin (top_of_set UNIV) S" using \closed S\ closed_closedin by fastforce then obtain g where "continuous_on UNIV g" "range g \ T" "\x. x \ S \ g x = f x" apply (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ c, where T=UNIV]) using \ANR T\ \S \ {}\ c homotopic_with_imp_subset1 apply fastforce+ done then show ?rhs by blast next assume ?rhs then obtain g where "continuous_on UNIV g" "range g \ T" "\x. x\S \ g x = f x" by blast then obtain c where "homotopic_with_canon (\h. True) UNIV T g (\x. c)" using nullhomotopic_from_contractible [of UNIV g T] contractible_UNIV by blast then have "homotopic_with_canon (\x. True) S T g (\x. c)" by (simp add: homotopic_from_subtopology) then show ?lhs by (force elim: homotopic_with_eq [of _ _ _ g "\x. c"] simp: \\x. x \ S \ g x = f x\) qed corollary\<^marker>\tag unimportant\ nullhomotopic_into_rel_frontier_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" and "convex T" "bounded T" and fim: "f ` S \ rel_frontier T" and "S \ {}" shows "(\c. homotopic_with_canon (\x. True) S (rel_frontier T) f (\x. c)) \ (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x \ S. g x = f x))" by (simp add: nullhomotopic_into_ANR_extension assms ANR_rel_frontier_convex) corollary\<^marker>\tag unimportant\ nullhomotopic_into_sphere_extension: fixes f :: "'a::euclidean_space \ 'b :: euclidean_space" assumes "closed S" and contf: "continuous_on S f" and "S \ {}" and fim: "f ` S \ sphere a r" shows "((\c. homotopic_with_canon (\x. True) S (sphere a r) f (\x. c)) \ (\g. continuous_on UNIV g \ range g \ sphere a r \ (\x \ S. g x = f x)))" (is "?lhs = ?rhs") proof (cases "r = 0") case True with fim show ?thesis apply auto using fim continuous_on_const apply fastforce by (metis contf contractible_sing nullhomotopic_into_contractible) next case False then have eq: "sphere a r = rel_frontier (cball a r)" by simp show ?thesis using fim unfolding eq apply (rule nullhomotopic_into_rel_frontier_extension [OF \closed S\ contf convex_cball bounded_cball]) apply (rule \S \ {}\) done qed proposition\<^marker>\tag unimportant\ Borsuk_map_essential_bounded_component: fixes a :: "'a :: euclidean_space" assumes "compact S" and "a \ S" shows "bounded (connected_component_set (- S) a) \ \(\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\x. c))" (is "?lhs = ?rhs") proof (cases "S = {}") case True then show ?thesis by simp next case False have "closed S" "bounded S" using \compact S\ compact_eq_bounded_closed by auto have s01: "(\x. (x - a) /\<^sub>R norm (x - a)) ` S \ sphere 0 1" using \a \ S\ by clarsimp (metis dist_eq_0_iff dist_norm mult.commute right_inverse) have aincc: "a \ connected_component_set (- S) a" by (simp add: \a \ S\) obtain r where "r>0" and r: "S \ ball 0 r" using bounded_subset_ballD \bounded S\ by blast have "\ ?rhs \ \ ?lhs" proof assume notr: "\ ?rhs" have nog: "\g. continuous_on (S \ connected_component_set (- S) a) g \ g ` (S \ connected_component_set (- S) a) \ sphere 0 1 \ (\x\S. g x = (x - a) /\<^sub>R norm (x - a))" if "bounded (connected_component_set (- S) a)" apply (rule non_extensible_Borsuk_map [OF \compact S\ componentsI _ aincc]) using \a \ S\ that by auto obtain g where "range g \ sphere 0 1" "continuous_on UNIV g" "\x. x \ S \ g x = (x - a) /\<^sub>R norm (x - a)" using notr by (auto simp: nullhomotopic_into_sphere_extension [OF \closed S\ continuous_on_Borsuk_map [OF \a \ S\] False s01]) with \a \ S\ show "\ ?lhs" apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog) apply (drule_tac x=g in spec) using continuous_on_subset by fastforce next assume "\ ?lhs" then obtain b where b: "b \ connected_component_set (- S) a" and "r \ norm b" using bounded_iff linear by blast then have bnot: "b \ ball 0 r" by simp have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. (x - b) /\<^sub>R norm (x - b))" apply (rule Borsuk_maps_homotopic_in_path_component) using \closed S\ b open_Compl open_path_connected_component apply fastforce done moreover obtain c where "homotopic_with_canon (\x. True) (ball 0 r) (sphere 0 1) (\x. inverse (norm (x - b)) *\<^sub>R (x - b)) (\x. c)" proof (rule nullhomotopic_from_contractible) show "contractible (ball (0::'a) r)" by (metis convex_imp_contractible convex_ball) show "continuous_on (ball 0 r) (\x. inverse(norm (x - b)) *\<^sub>R (x - b))" by (rule continuous_on_Borsuk_map [OF bnot]) show "(\x. (x - b) /\<^sub>R norm (x - b)) ` ball 0 r \ sphere 0 1" using bnot Borsuk_map_into_sphere by blast qed blast ultimately have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. c)" by (meson homotopic_with_subset_left homotopic_with_trans r) then show "\ ?rhs" by blast qed then show ?thesis by blast qed lemma homotopic_Borsuk_maps_in_bounded_component: fixes a :: "'a :: euclidean_space" assumes "compact S" and "a \ S"and "b \ S" and boc: "bounded (connected_component_set (- S) a)" and hom: "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. (x - b) /\<^sub>R norm (x - b))" shows "connected_component (- S) a b" proof (rule ccontr) assume notcc: "\ connected_component (- S) a b" let ?T = "S \ connected_component_set (- S) a" have "\g. continuous_on (S \ connected_component_set (- S) a) g \ g ` (S \ connected_component_set (- S) a) \ sphere 0 1 \ (\x\S. g x = (x - a) /\<^sub>R norm (x - a))" by (simp add: \a \ S\ componentsI non_extensible_Borsuk_map [OF \compact S\ _ boc]) moreover obtain g where "continuous_on (S \ connected_component_set (- S) a) g" "g ` (S \ connected_component_set (- S) a) \ sphere 0 1" "\x. x \ S \ g x = (x - a) /\<^sub>R norm (x - a)" proof (rule Borsuk_homotopy_extension_homotopic) show "closedin (top_of_set ?T) S" by (simp add: \compact S\ closed_subset compact_imp_closed) show "continuous_on ?T (\x. (x - b) /\<^sub>R norm (x - b))" by (simp add: \b \ S\ notcc continuous_on_Borsuk_map) show "(\x. (x - b) /\<^sub>R norm (x - b)) ` ?T \ sphere 0 1" by (simp add: \b \ S\ notcc Borsuk_map_into_sphere) show "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - b) /\<^sub>R norm (x - b)) (\x. (x - a) /\<^sub>R norm (x - a))" by (simp add: hom homotopic_with_symD) qed (auto simp: ANR_sphere intro: that) ultimately show False by blast qed lemma Borsuk_maps_homotopic_in_connected_component_eq: fixes a :: "'a :: euclidean_space" assumes S: "compact S" "a \ S" "b \ S" and 2: "2 \ DIM('a)" shows "(homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. (x - b) /\<^sub>R norm (x - b)) \ connected_component (- S) a b)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (cases "bounded(connected_component_set (- S) a)") case True show ?thesis by (rule homotopic_Borsuk_maps_in_bounded_component [OF S True L]) next case not_bo_a: False show ?thesis proof (cases "bounded(connected_component_set (- S) b)") case True show ?thesis using homotopic_Borsuk_maps_in_bounded_component [OF S] by (simp add: L True assms connected_component_sym homotopic_Borsuk_maps_in_bounded_component homotopic_with_sym) next case False then show ?thesis using cobounded_unique_unbounded_component [of "-S" a b] \compact S\ not_bo_a by (auto simp: compact_eq_bounded_closed assms connected_component_eq_eq) qed qed next assume R: ?rhs then have "path_component (- S) a b" using assms(1) compact_eq_bounded_closed open_Compl open_path_connected_component_set by fastforce then show ?lhs by (simp add: Borsuk_maps_homotopic_in_path_component) qed subsection\More extension theorems\ lemma extension_from_clopen: assumes ope: "openin (top_of_set S) T" and clo: "closedin (top_of_set S) T" and contf: "continuous_on T f" and fim: "f ` T \ U" and null: "U = {} \ S = {}" obtains g where "continuous_on S g" "g ` S \ U" "\x. x \ T \ g x = f x" proof (cases "U = {}") case True then show ?thesis by (simp add: null that) next case False then obtain a where "a \ U" by auto let ?g = "\x. if x \ T then f x else a" have Seq: "S = T \ (S - T)" using clo closedin_imp_subset by fastforce show ?thesis proof have "continuous_on (T \ (S - T)) ?g" apply (rule continuous_on_cases_local) using Seq clo ope by (auto simp: contf intro: continuous_on_cases_local) with Seq show "continuous_on S ?g" by metis show "?g ` S \ U" using \a \ U\ fim by auto show "\x. x \ T \ ?g x = f x" by auto qed qed lemma extension_from_component: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes S: "locally connected S \ compact S" and "ANR U" and C: "C \ components S" and contf: "continuous_on C f" and fim: "f ` C \ U" obtains g where "continuous_on S g" "g ` S \ U" "\x. x \ C \ g x = f x" proof - obtain T g where ope: "openin (top_of_set S) T" and clo: "closedin (top_of_set S) T" and "C \ T" and contg: "continuous_on T g" and gim: "g ` T \ U" and gf: "\x. x \ C \ g x = f x" using S proof assume "locally connected S" show ?thesis by (metis C \locally connected S\ openin_components_locally_connected closedin_component contf fim order_refl that) next assume "compact S" then obtain W g where "C \ W" and opeW: "openin (top_of_set S) W" and contg: "continuous_on W g" and gim: "g ` W \ U" and gf: "\x. x \ C \ g x = f x" using ANR_imp_absolute_neighbourhood_extensor [of U C f S] C \ANR U\ closedin_component contf fim by blast then obtain V where "open V" and V: "W = S \ V" by (auto simp: openin_open) moreover have "locally compact S" by (simp add: \compact S\ closed_imp_locally_compact compact_imp_closed) ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \ K" "K \ V" by (metis C Int_subset_iff \C \ W\ \compact S\ compact_components Sura_Bura_clopen_subset) show ?thesis proof show "closedin (top_of_set S) K" by (meson \compact K\ \compact S\ closedin_compact_eq opeK openin_imp_subset) show "continuous_on K g" by (metis Int_subset_iff V \K \ V\ contg continuous_on_subset opeK openin_subtopology subset_eq) show "g ` K \ U" using V \K \ V\ gim opeK openin_imp_subset by fastforce qed (use opeK gf \C \ K\ in auto) qed obtain h where "continuous_on S h" "h ` S \ U" "\x. x \ T \ h x = g x" using extension_from_clopen by (metis C bot.extremum_uniqueI clo contg gim fim image_is_empty in_components_nonempty ope) then show ?thesis by (metis \C \ T\ gf subset_eq that) qed lemma tube_lemma: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and S: "S \ {}" "(\x. (x,a)) ` S \ U" and ope: "openin (top_of_set (S \ T)) U" obtains V where "openin (top_of_set T) V" "a \ V" "S \ V \ U" proof - let ?W = "{y. \x. x \ S \ (x, y) \ (S \ T - U)}" have "U \ S \ T" "closedin (top_of_set (S \ T)) (S \ T - U)" using ope by (auto simp: openin_closedin_eq) then have "closedin (top_of_set T) ?W" using \compact S\ closedin_compact_projection by blast moreover have "a \ T - ?W" using \U \ S \ T\ S by auto moreover have "S \ (T - ?W) \ U" by auto ultimately show ?thesis by (metis (no_types, lifting) Sigma_cong closedin_def that topspace_euclidean_subtopology) qed lemma tube_lemma_gen: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "S \ {}" "T \ T'" "S \ T \ U" and ope: "openin (top_of_set (S \ T')) U" obtains V where "openin (top_of_set T') V" "T \ V" "S \ V \ U" proof - have "\x. x \ T \ \V. openin (top_of_set T') V \ x \ V \ S \ V \ U" using assms by (auto intro: tube_lemma [OF \compact S\]) then obtain F where F: "\x. x \ T \ openin (top_of_set T') (F x) \ x \ F x \ S \ F x \ U" by metis show ?thesis proof show "openin (top_of_set T') (\(F ` T))" using F by blast show "T \ \(F ` T)" using F by blast show "S \ \(F ` T) \ U" using F by auto qed qed proposition\<^marker>\tag unimportant\ homotopic_neighbourhood_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and fim: "f ` S \ U" and contg: "continuous_on S g" and gim: "g ` S \ U" and clo: "closedin (top_of_set S) T" and "ANR U" and hom: "homotopic_with_canon (\x. True) T U f g" obtains V where "T \ V" "openin (top_of_set S) V" "homotopic_with_canon (\x. True) V U f g" proof - have "T \ S" using clo closedin_imp_subset by blast obtain h where conth: "continuous_on ({0..1::real} \ T) h" and him: "h ` ({0..1} \ T) \ U" and h0: "\x. h(0, x) = f x" and h1: "\x. h(1, x) = g x" using hom by (auto simp: homotopic_with_def) define h' where "h' \ \z. if fst z \ {0} then f(snd z) else if fst z \ {1} then g(snd z) else h z" let ?S0 = "{0::real} \ S" and ?S1 = "{1::real} \ S" have "continuous_on(?S0 \ (?S1 \ {0..1} \ T)) h'" unfolding h'_def proof (intro continuous_on_cases_local) show "closedin (top_of_set (?S0 \ (?S1 \ {0..1} \ T))) ?S0" "closedin (top_of_set (?S1 \ {0..1} \ T)) ?S1" using \T \ S\ by (force intro: closedin_Times closedin_subset_trans [of "{0..1} \ S"])+ show "closedin (top_of_set (?S0 \ (?S1 \ {0..1} \ T))) (?S1 \ {0..1} \ T)" "closedin (top_of_set (?S1 \ {0..1} \ T)) ({0..1} \ T)" using \T \ S\ by (force intro: clo closedin_Times closedin_subset_trans [of "{0..1} \ S"])+ show "continuous_on (?S0) (\x. f (snd x))" by (intro continuous_intros continuous_on_compose2 [OF contf]) auto show "continuous_on (?S1) (\x. g (snd x))" by (intro continuous_intros continuous_on_compose2 [OF contg]) auto qed (use h0 h1 conth in auto) then have "continuous_on ({0,1} \ S \ ({0..1} \ T)) h'" by (metis Sigma_Un_distrib1 Un_assoc insert_is_Un) moreover have "h' ` ({0,1} \ S \ {0..1} \ T) \ U" using fim gim him \T \ S\ unfolding h'_def by force moreover have "closedin (top_of_set ({0..1::real} \ S)) ({0,1} \ S \ {0..1::real} \ T)" by (intro closedin_Times closedin_Un clo) (simp_all add: closed_subset) ultimately obtain W k where W: "({0,1} \ S) \ ({0..1} \ T) \ W" and opeW: "openin (top_of_set ({0..1} \ S)) W" and contk: "continuous_on W k" and kim: "k ` W \ U" and kh': "\x. x \ ({0,1} \ S) \ ({0..1} \ T) \ k x = h' x" by (metis ANR_imp_absolute_neighbourhood_extensor [OF \ANR U\, of "({0,1} \ S) \ ({0..1} \ T)" h' "{0..1} \ S"]) obtain T' where opeT': "openin (top_of_set S) T'" and "T \ T'" and TW: "{0..1} \ T' \ W" using tube_lemma_gen [of "{0..1::real}" T S W] W \T \ S\ opeW by auto moreover have "homotopic_with_canon (\x. True) T' U f g" proof (simp add: homotopic_with, intro exI conjI) show "continuous_on ({0..1} \ T') k" using TW continuous_on_subset contk by auto show "k ` ({0..1} \ T') \ U" using TW kim by fastforce have "T' \ S" by (meson opeT' subsetD openin_imp_subset) then show "\x\T'. k (0, x) = f x" "\x\T'. k (1, x) = g x" by (auto simp: kh' h'_def) qed ultimately show ?thesis by (blast intro: that) qed text\ Homotopy on a union of closed-open sets.\ proposition\<^marker>\tag unimportant\ homotopic_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S" and "\S. S \ \ \ homotopic_with_canon (\x. True) S T f g" shows "homotopic_with_canon (\x. True) (\\) T f g" proof - obtain \ where "\ \ \" "countable \" and eqU: "\\ = \\" using Lindelof_openin assms by blast show ?thesis proof (cases "\ = {}") case True then show ?thesis by (metis Union_empty eqU homotopic_with_canon_on_empty) next case False then obtain V :: "nat \ 'a set" where V: "range V = \" using range_from_nat_into \countable \\ by metis with \\ \ \\ have clo: "\n. closedin (top_of_set (\\)) (V n)" and ope: "\n. openin (top_of_set (\\)) (V n)" and hom: "\n. homotopic_with_canon (\x. True) (V n) T f g" using assms by auto then obtain h where conth: "\n. continuous_on ({0..1::real} \ V n) (h n)" and him: "\n. h n ` ({0..1} \ V n) \ T" and h0: "\n. \x. x \ V n \ h n (0, x) = f x" and h1: "\n. \x. x \ V n \ h n (1, x) = g x" by (simp add: homotopic_with) metis have wop: "b \ V x \ \k. b \ V k \ (\j V j)" for b x using nat_less_induct [where P = "\i. b \ V i"] by meson obtain \ where cont: "continuous_on ({0..1} \ \(V ` UNIV)) \" and eq: "\x i. \x \ {0..1} \ \(V ` UNIV) \ {0..1} \ (V i - (\m \ \ x = h i x" proof (rule pasting_lemma_exists) let ?X = "top_of_set ({0..1::real} \ \(range V))" show "topspace ?X \ (\i. {0..1::real} \ (V i - (\m \(V ` UNIV))) ({0..1::real} \ (V i - (\m(V ` UNIV))) (V i)" using ope V eqU by auto show "closedin (top_of_set (\(V ` UNIV))) (\m (V i - \ (V ` {..i j x. x \ topspace ?X \ {0..1} \ (V i - (\m {0..1} \ (V j - (\m h i x = h j x" by clarsimp (metis lessThan_iff linorder_neqE_nat) qed auto show ?thesis proof (simp add: homotopic_with eqU [symmetric], intro exI conjI ballI) show "continuous_on ({0..1} \ \\) \" using V eqU by (blast intro!: continuous_on_subset [OF cont]) show "\` ({0..1} \ \\) \ T" proof clarsimp fix t :: real and y :: "'a" and X :: "'a set" assume "y \ X" "X \ \" and t: "0 \ t" "t \ 1" then obtain k where "y \ V k" and j: "\j V j" by (metis image_iff V wop) with him t show "\(t, y) \ T" by (subst eq) force+ qed fix X y assume "X \ \" "y \ X" then obtain k where "y \ V k" and j: "\j V j" by (metis image_iff V wop) then show "\(0, y) = f y" and "\(1, y) = g y" by (subst eq [where i=k]; force simp: h0 h1)+ qed qed qed lemma homotopic_on_components_eq: fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set" assumes S: "locally connected S \ compact S" and "ANR T" shows "homotopic_with_canon (\x. True) S T f g \ (continuous_on S f \ f ` S \ T \ continuous_on S g \ g ` S \ T) \ (\C \ components S. homotopic_with_canon (\x. True) C T f g)" (is "?lhs \ ?C \ ?rhs") proof - have "continuous_on S f" "f ` S \ T" "continuous_on S g" "g ` S \ T" if ?lhs using homotopic_with_imp_continuous homotopic_with_imp_subset1 homotopic_with_imp_subset2 that by blast+ moreover have "?lhs \ ?rhs" if contf: "continuous_on S f" and fim: "f ` S \ T" and contg: "continuous_on S g" and gim: "g ` S \ T" proof assume ?lhs with that show ?rhs by (simp add: homotopic_with_subset_left in_components_subset) next assume R: ?rhs have "\U. C \ U \ closedin (top_of_set S) U \ openin (top_of_set S) U \ homotopic_with_canon (\x. True) U T f g" if C: "C \ components S" for C proof - have "C \ S" by (simp add: in_components_subset that) show ?thesis using S proof assume "locally connected S" show ?thesis proof (intro exI conjI) show "closedin (top_of_set S) C" by (simp add: closedin_component that) show "openin (top_of_set S) C" by (simp add: \locally connected S\ openin_components_locally_connected that) show "homotopic_with_canon (\x. True) C T f g" by (simp add: R that) qed auto next assume "compact S" have hom: "homotopic_with_canon (\x. True) C T f g" using R that by blast obtain U where "C \ U" and opeU: "openin (top_of_set S) U" and hom: "homotopic_with_canon (\x. True) U T f g" using homotopic_neighbourhood_extension [OF contf fim contg gim _ \ANR T\ hom] \C \ components S\ closedin_component by blast then obtain V where "open V" and V: "U = S \ V" by (auto simp: openin_open) moreover have "locally compact S" by (simp add: \compact S\ closed_imp_locally_compact compact_imp_closed) ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \ K" "K \ V" by (metis C Int_subset_iff Sura_Bura_clopen_subset \C \ U\ \compact S\ compact_components) show ?thesis proof (intro exI conjI) show "closedin (top_of_set S) K" by (meson \compact K\ \compact S\ closedin_compact_eq opeK openin_imp_subset) show "homotopic_with_canon (\x. True) K T f g" using V \K \ V\ hom homotopic_with_subset_left opeK openin_imp_subset by fastforce qed (use opeK \C \ K\ in auto) qed qed then obtain \ where \: "\C. C \ components S \ C \ \ C" and clo\: "\C. C \ components S \ closedin (top_of_set S) (\ C)" and ope\: "\C. C \ components S \ openin (top_of_set S) (\ C)" and hom\: "\C. C \ components S \ homotopic_with_canon (\x. True) (\ C) T f g" by metis have Seq: "S = \ (\ ` components S)" proof show "S \ \ (\ ` components S)" by (metis Sup_mono Union_components \ imageI) show "\ (\ ` components S) \ S" using ope\ openin_imp_subset by fastforce qed show ?lhs apply (subst Seq) apply (rule homotopic_on_clopen_Union) using Seq clo\ ope\ hom\ by auto qed ultimately show ?thesis by blast qed lemma cohomotopically_trivial_on_components: fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set" assumes S: "locally connected S \ compact S" and "ANR T" shows "(\f g. continuous_on S f \ f ` S \ T \ continuous_on S g \ g ` S \ T \ homotopic_with_canon (\x. True) S T f g) \ (\C\components S. \f g. continuous_on C f \ f ` C \ T \ continuous_on C g \ g ` C \ T \ homotopic_with_canon (\x. True) C T f g)" (is "?lhs = ?rhs") proof assume L[rule_format]: ?lhs show ?rhs proof clarify fix C f g assume contf: "continuous_on C f" and fim: "f ` C \ T" and contg: "continuous_on C g" and gim: "g ` C \ T" and C: "C \ components S" obtain f' where contf': "continuous_on S f'" and f'im: "f' ` S \ T" and f'f: "\x. x \ C \ f' x = f x" using extension_from_component [OF S \ANR T\ C contf fim] by metis obtain g' where contg': "continuous_on S g'" and g'im: "g' ` S \ T" and g'g: "\x. x \ C \ g' x = g x" using extension_from_component [OF S \ANR T\ C contg gim] by metis have "homotopic_with_canon (\x. True) C T f' g'" using L [OF contf' f'im contg' g'im] homotopic_with_subset_left C in_components_subset by fastforce then show "homotopic_with_canon (\x. True) C T f g" using f'f g'g homotopic_with_eq by force qed next assume R [rule_format]: ?rhs show ?lhs proof clarify fix f g assume contf: "continuous_on S f" and fim: "f ` S \ T" and contg: "continuous_on S g" and gim: "g ` S \ T" moreover have "homotopic_with_canon (\x. True) C T f g" if "C \ components S" for C using R [OF that] by (meson contf contg continuous_on_subset fim gim image_mono in_components_subset order.trans that) ultimately show "homotopic_with_canon (\x. True) S T f g" by (subst homotopic_on_components_eq [OF S \ANR T\]) auto qed qed subsection\The complement of a set and path-connectedness\ text\Complement in dimension N > 1 of set homeomorphic to any interval in any dimension is (path-)connected. This naively generalizes the argument in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer fixed point theorem", American Mathematical Monthly 1984.\ lemma unbounded_components_complement_absolute_retract: fixes S :: "'a::euclidean_space set" assumes C: "C \ components(- S)" and S: "compact S" "AR S" shows "\ bounded C" proof - obtain y where y: "C = connected_component_set (- S) y" and "y \ S" using C by (auto simp: components_def) have "open(- S)" using S by (simp add: closed_open compact_eq_bounded_closed) have "S retract_of UNIV" using S compact_AR by blast then obtain r where contr: "continuous_on UNIV r" and ontor: "range r \ S" and r: "\x. x \ S \ r x = x" by (auto simp: retract_of_def retraction_def) show ?thesis proof assume "bounded C" have "connected_component_set (- S) y \ S" proof (rule frontier_subset_retraction) show "bounded (connected_component_set (- S) y)" using \bounded C\ y by blast show "frontier (connected_component_set (- S) y) \ S" using C \compact S\ compact_eq_bounded_closed frontier_of_components_closed_complement y by blast show "continuous_on (closure (connected_component_set (- S) y)) r" by (blast intro: continuous_on_subset [OF contr]) qed (use ontor r in auto) with \y \ S\ show False by force qed qed lemma connected_complement_absolute_retract: fixes S :: "'a::euclidean_space set" assumes S: "compact S" "AR S" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S retract_of UNIV" using S compact_AR by blast show ?thesis apply (clarsimp simp: connected_iff_connected_component_eq) apply (rule cobounded_unique_unbounded_component [OF _ 2]) apply (simp add: \compact S\ compact_imp_bounded) apply (meson ComplI S componentsI unbounded_components_complement_absolute_retract)+ done qed lemma path_connected_complement_absolute_retract: fixes S :: "'a::euclidean_space set" assumes "compact S" "AR S" "2 \ DIM('a)" shows "path_connected(- S)" using connected_complement_absolute_retract [OF assms] using \compact S\ compact_eq_bounded_closed connected_open_path_connected by blast theorem connected_complement_homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes hom: "S homeomorphic T" and T: "convex T" "compact T" and 2: "2 \ DIM('a)" shows "connected(- S)" proof (cases "S = {}") case True then show ?thesis by (simp add: connected_UNIV) next case False show ?thesis proof (rule connected_complement_absolute_retract) show "compact S" using \compact T\ hom homeomorphic_compactness by auto show "AR S" by (meson AR_ANR False \convex T\ convex_imp_ANR convex_imp_contractible hom homeomorphic_ANR_iff_ANR homeomorphic_contractible_eq) qed (rule 2) qed corollary path_connected_complement_homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes hom: "S homeomorphic T" "convex T" "compact T" "2 \ DIM('a)" shows "path_connected(- S)" using connected_complement_homeomorphic_convex_compact [OF assms] using \compact T\ compact_eq_bounded_closed connected_open_path_connected hom homeomorphic_compactness by blast lemma path_connected_complement_homeomorphic_interval: fixes S :: "'a::euclidean_space set" assumes "S homeomorphic cbox a b" "2 \ DIM('a)" shows "path_connected(-S)" using assms compact_cbox convex_box(1) path_connected_complement_homeomorphic_convex_compact by blast lemma connected_complement_homeomorphic_interval: fixes S :: "'a::euclidean_space set" assumes "S homeomorphic cbox a b" "2 \ DIM('a)" shows "connected(-S)" using assms path_connected_complement_homeomorphic_interval path_connected_imp_connected by blast + +lemma path_connected_arc_complement: + fixes \ :: "real \ 'a::euclidean_space" + assumes "arc \" "2 \ DIM('a)" + shows "path_connected(- path_image \)" +proof - + have "path_image \ homeomorphic {0..1::real}" + by (simp add: assms homeomorphic_arc_image_interval) + then + show ?thesis + apply (rule path_connected_complement_homeomorphic_convex_compact) + apply (auto simp: assms) + done +qed + +lemma connected_arc_complement: + fixes \ :: "real \ 'a::euclidean_space" + assumes "arc \" "2 \ DIM('a)" + shows "connected(- path_image \)" + by (simp add: assms path_connected_arc_complement path_connected_imp_connected) + +lemma inside_arc_empty: + fixes \ :: "real \ 'a::euclidean_space" + assumes "arc \" + shows "inside(path_image \) = {}" +proof (cases "DIM('a) = 1") + case True + then show ?thesis + using assms connected_arc_image connected_convex_1_gen inside_convex by blast +next + case False + show ?thesis + proof (rule inside_bounded_complement_connected_empty) + show "connected (- path_image \)" + apply (rule connected_arc_complement [OF assms]) + using False + by (metis DIM_ge_Suc0 One_nat_def Suc_1 not_less_eq_eq order_class.order.antisym) + show "bounded (path_image \)" + by (simp add: assms bounded_arc_image) + qed +qed + +lemma inside_simple_curve_imp_closed: + fixes \ :: "real \ 'a::euclidean_space" + shows "\simple_path \; x \ inside(path_image \)\ \ pathfinish \ = pathstart \" + using arc_simple_path inside_arc_empty by blast + end diff --git a/src/HOL/Analysis/Winding_Numbers.thy b/src/HOL/Analysis/Winding_Numbers.thy --- a/src/HOL/Analysis/Winding_Numbers.thy +++ b/src/HOL/Analysis/Winding_Numbers.thy @@ -1,1211 +1,1330 @@ section \Winding Numbers\ - -text\By John Harrison et al. Ported from HOL Light by L C Paulson (2017)\ - -theory Winding_Numbers -imports - Polytope - Jordan_Curve - Riemann_Mapping +theory Winding_Numbers + imports Cauchy_Integral_Theorem begin -lemma simply_connected_inside_simple_path: - fixes p :: "real \ complex" - shows "simple_path p \ simply_connected(inside(path_image p))" - using Jordan_inside_outside connected_simple_path_image inside_simple_curve_imp_closed simply_connected_eq_frontier_properties - by fastforce +text\We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\ -lemma simply_connected_Int: - fixes S :: "complex set" - assumes "open S" "open T" "simply_connected S" "simply_connected T" "connected (S \ T)" - shows "simply_connected (S \ T)" - using assms by (force simp: simply_connected_eq_winding_number_zero open_Int) +subsection \Basic Winding Numbers\ -subsection\Winding number for a triangle\ +definition\<^marker>\tag important\ winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where + "winding_number_prop \ z e p n \ + valid_path p \ z \ path_image p \ + pathstart p = pathstart \ \ + pathfinish p = pathfinish \ \ + (\t \ {0..1}. norm(\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" -lemma wn_triangle1: - assumes "0 \ interior(convex hull {a,b,c})" - shows "\ (Im(a/b) \ 0 \ 0 \ Im(b/c))" +definition\<^marker>\tag important\ winding_number:: "[real \ complex, complex] \ complex" where + "winding_number \ z \ SOME n. \e > 0. \p. winding_number_prop \ z e p n" + +lemma winding_number: + assumes "path \" "z \ path_image \" "0 < e" + shows "\p. winding_number_prop \ z e p (winding_number \ z)" proof - - { assume 0: "Im(a/b) \ 0" "0 \ Im(b/c)" - have "0 \ interior (convex hull {a,b,c})" - proof (cases "a=0 \ b=0 \ c=0") - case True then show ?thesis - by (auto simp: not_in_interior_convex_hull_3) - next - case False - then have "b \ 0" by blast - { fix x y::complex and u::real - assume eq_f': "Im x * Re b \ Im b * Re x" "Im y * Re b \ Im b * Re y" "0 \ u" "u \ 1" - then have "((1 - u) * Im x) * Re b \ Im b * ((1 - u) * Re x)" - by (simp add: mult_left_mono mult.assoc mult.left_commute [of "Im b"]) - then have "((1 - u) * Im x + u * Im y) * Re b \ Im b * ((1 - u) * Re x + u * Re y)" - using eq_f' ordered_comm_semiring_class.comm_mult_left_mono - by (fastforce simp add: algebra_simps) - } - with False 0 have "convex hull {a,b,c} \ {z. Im z * Re b \ Im b * Re z}" - apply (simp add: Complex.Im_divide divide_simps complex_neq_0 [symmetric]) - apply (simp add: algebra_simps) - apply (rule hull_minimal) - apply (auto simp: algebra_simps convex_alt) + have "path_image \ \ UNIV - {z}" + using assms by blast + then obtain d + where d: "d>0" + and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ + (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ + pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ + path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ + (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" + using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) + then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ + (\t \ {0..1}. norm(h t - \ t) < d/2)" + using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto + define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" + have "\n. \e > 0. \p. winding_number_prop \ z e p n" + proof (rule_tac x=nn in exI, clarify) + fix e::real + assume e: "e>0" + obtain p where p: "polynomial_function p \ + pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d/2))" + using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by auto + have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" + by (auto simp: intro!: holomorphic_intros) + then show "\p. winding_number_prop \ z e p nn" + apply (rule_tac x=p in exI) + using pi_eq [of h p] h p d + apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def winding_number_prop_def) done - moreover have "0 \ interior({z. Im z * Re b \ Im b * Re z})" - proof - assume "0 \ interior {z. Im z * Re b \ Im b * Re z}" - then obtain e where "e>0" and e: "ball 0 e \ {z. Im z * Re b \ Im b * Re z}" - by (meson mem_interior) - define z where "z \ - sgn (Im b) * (e/3) + sgn (Re b) * (e/3) * \" - have "z \ ball 0 e" - using \e>0\ - apply (simp add: z_def dist_norm) - apply (rule le_less_trans [OF norm_triangle_ineq4]) - apply (simp add: norm_mult abs_sgn_eq) - done - then have "z \ {z. Im z * Re b \ Im b * Re z}" - using e by blast - then show False - using \e>0\ \b \ 0\ - apply (simp add: z_def dist_norm sgn_if less_eq_real_def mult_less_0_iff complex.expand split: if_split_asm) - apply (auto simp: algebra_simps) - apply (meson less_asym less_trans mult_pos_pos neg_less_0_iff_less) - by (metis less_asym mult_pos_pos neg_less_0_iff_less) - qed - ultimately show ?thesis - using interior_mono by blast qed - } with assms show ?thesis by blast + then show ?thesis + unfolding winding_number_def by (rule someI2_ex) (blast intro: \0) qed -lemma wn_triangle2_0: - assumes "0 \ interior(convex hull {a,b,c})" - shows - "0 < Im((b - a) * cnj (b)) \ - 0 < Im((c - b) * cnj (c)) \ - 0 < Im((a - c) * cnj (a)) - \ - Im((b - a) * cnj (b)) < 0 \ - 0 < Im((b - c) * cnj (b)) \ - 0 < Im((a - b) * cnj (a)) \ - 0 < Im((c - a) * cnj (c))" +lemma winding_number_unique: + assumes \: "path \" "z \ path_image \" + and pi: "\e. e>0 \ \p. winding_number_prop \ z e p n" + shows "winding_number \ z = n" proof - - have [simp]: "{b,c,a} = {a,b,c}" "{c,a,b} = {a,b,c}" by auto - show ?thesis - using wn_triangle1 [OF assms] wn_triangle1 [of b c a] wn_triangle1 [of c a b] assms - by (auto simp: algebra_simps Im_complex_div_gt_0 Im_complex_div_lt_0 not_le not_less) -qed - -lemma wn_triangle2: - assumes "z \ interior(convex hull {a,b,c})" - shows "0 < Im((b - a) * cnj (b - z)) \ - 0 < Im((c - b) * cnj (c - z)) \ - 0 < Im((a - c) * cnj (a - z)) - \ - Im((b - a) * cnj (b - z)) < 0 \ - 0 < Im((b - c) * cnj (b - z)) \ - 0 < Im((a - b) * cnj (a - z)) \ - 0 < Im((c - a) * cnj (c - z))" -proof - - have 0: "0 \ interior(convex hull {a-z, b-z, c-z})" - using assms convex_hull_translation [of "-z" "{a,b,c}"] - interior_translation [of "-z"] - by (simp cong: image_cong_simp) - show ?thesis using wn_triangle2_0 [OF 0] + have "path_image \ \ UNIV - {z}" + using assms by blast + then obtain e + where e: "e>0" + and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; + (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); + pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ + contour_integral h2 f = contour_integral h1 f" + using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) + obtain p where p: "winding_number_prop \ z e p n" + using pi [OF e] by blast + obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" + using winding_number [OF \ e] by blast + have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" + using p by (auto simp: winding_number_prop_def) + also have "\ = contour_integral q (\w. 1 / (w - z))" + proof (rule pi_eq) + show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" + by (auto intro!: holomorphic_intros) + qed (use p q in \auto simp: winding_number_prop_def norm_minus_commute\) + also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" + using q by (auto simp: winding_number_prop_def) + finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . + then show ?thesis by simp qed -lemma wn_triangle3: - assumes z: "z \ interior(convex hull {a,b,c})" - and "0 < Im((b-a) * cnj (b-z))" - "0 < Im((c-b) * cnj (c-z))" - "0 < Im((a-c) * cnj (a-z))" - shows "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = 1" +(*NB not winding_number_prop here due to the loop in p*) +lemma winding_number_unique_loop: + assumes \: "path \" "z \ path_image \" + and loop: "pathfinish \ = pathstart \" + and pi: + "\e. e>0 \ \p. valid_path p \ z \ path_image p \ + pathfinish p = pathstart p \ + (\t \ {0..1}. norm (\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" + shows "winding_number \ z = n" proof - - have znot[simp]: "z \ closed_segment a b" "z \ closed_segment b c" "z \ closed_segment c a" - using z interior_of_triangle [of a b c] - by (auto simp: closed_segment_def) - have gt0: "0 < Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z)" - using assms - by (simp add: winding_number_linepath_pos_lt path_image_join winding_number_join_pos_combined) - have lt2: "Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z) < 2" - using winding_number_lt_half_linepath [of _ a b] - using winding_number_lt_half_linepath [of _ b c] - using winding_number_lt_half_linepath [of _ c a] znot - apply (fastforce simp add: winding_number_join path_image_join) - done - show ?thesis - by (rule winding_number_eq_1) (simp_all add: path_image_join gt0 lt2) + have "path_image \ \ UNIV - {z}" + using assms by blast + then obtain e + where e: "e>0" + and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; + (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); + pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ + contour_integral h2 f = contour_integral h1 f" + using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) + obtain p where p: + "valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ + (\t \ {0..1}. norm (\ t - p t) < e) \ + contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" + using pi [OF e] by blast + obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" + using winding_number [OF \ e] by blast + have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" + using p by auto + also have "\ = contour_integral q (\w. 1 / (w - z))" + proof (rule pi_eq) + show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" + by (auto intro!: holomorphic_intros) + qed (use p q loop in \auto simp: winding_number_prop_def norm_minus_commute\) + also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" + using q by (auto simp: winding_number_prop_def) + finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . + then show ?thesis + by simp qed -proposition winding_number_triangle: - assumes z: "z \ interior(convex hull {a,b,c})" - shows "winding_number(linepath a b +++ linepath b c +++ linepath c a) z = - (if 0 < Im((b - a) * cnj (b - z)) then 1 else -1)" +proposition winding_number_valid_path: + assumes "valid_path \" "z \ path_image \" + shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" + by (rule winding_number_unique) + (use assms in \auto simp: valid_path_imp_path winding_number_prop_def\) + +proposition has_contour_integral_winding_number: + assumes \: "valid_path \" "z \ path_image \" + shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" +by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) + +lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" + by (simp add: winding_number_valid_path) + +lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" + by (simp add: path_image_subpath winding_number_valid_path) + +lemma winding_number_join: + assumes \1: "path \1" "z \ path_image \1" + and \2: "path \2" "z \ path_image \2" + and "pathfinish \1 = pathstart \2" + shows "winding_number(\1 +++ \2) z = winding_number \1 z + winding_number \2 z" +proof (rule winding_number_unique) + show "\p. winding_number_prop (\1 +++ \2) z e p + (winding_number \1 z + winding_number \2 z)" if "e > 0" for e + proof - + obtain p1 where "winding_number_prop \1 z e p1 (winding_number \1 z)" + using \0 < e\ \1 winding_number by blast + moreover + obtain p2 where "winding_number_prop \2 z e p2 (winding_number \2 z)" + using \0 < e\ \2 winding_number by blast + ultimately + have "winding_number_prop (\1+++\2) z e (p1+++p2) (winding_number \1 z + winding_number \2 z)" + using assms + apply (simp add: winding_number_prop_def not_in_path_image_join contour_integrable_inversediff algebra_simps) + apply (auto simp: joinpaths_def) + done + then show ?thesis + by blast + qed +qed (use assms in \auto simp: not_in_path_image_join\) + +lemma winding_number_reversepath: + assumes "path \" "z \ path_image \" + shows "winding_number(reversepath \) z = - (winding_number \ z)" +proof (rule winding_number_unique) + show "\p. winding_number_prop (reversepath \) z e p (- winding_number \ z)" if "e > 0" for e + proof - + obtain p where "winding_number_prop \ z e p (winding_number \ z)" + using \0 < e\ assms winding_number by blast + then have "winding_number_prop (reversepath \) z e (reversepath p) (- winding_number \ z)" + using assms + apply (simp add: winding_number_prop_def contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) + apply (auto simp: reversepath_def) + done + then show ?thesis + by blast + qed +qed (use assms in auto) + +lemma winding_number_shiftpath: + assumes \: "path \" "z \ path_image \" + and "pathfinish \ = pathstart \" "a \ {0..1}" + shows "winding_number(shiftpath a \) z = winding_number \ z" +proof (rule winding_number_unique_loop) + show "\p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ + (\t\{0..1}. cmod (shiftpath a \ t - p t) < e) \ + contour_integral p (\w. 1 / (w - z)) = + complex_of_real (2 * pi) * \ * winding_number \ z" + if "e > 0" for e + proof - + obtain p where "winding_number_prop \ z e p (winding_number \ z)" + using \0 < e\ assms winding_number by blast + then show ?thesis + apply (rule_tac x="shiftpath a p" in exI) + using assms that + apply (auto simp: winding_number_prop_def path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath contour_integral_shiftpath) + apply (simp add: shiftpath_def) + done + qed +qed (use assms in \auto simp: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath\) + +lemma winding_number_split_linepath: + assumes "c \ closed_segment a b" "z \ closed_segment a b" + shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" proof - - have [simp]: "{a,c,b} = {a,b,c}" by auto - have znot[simp]: "z \ closed_segment a b" "z \ closed_segment b c" "z \ closed_segment c a" - using z interior_of_triangle [of a b c] - by (auto simp: closed_segment_def) - then have [simp]: "z \ closed_segment b a" "z \ closed_segment c b" "z \ closed_segment a c" - using closed_segment_commute by blast+ - have *: "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = - winding_number (reversepath (linepath a c +++ linepath c b +++ linepath b a)) z" - by (simp add: reversepath_joinpaths winding_number_join not_in_path_image_join) - show ?thesis - using wn_triangle2 [OF z] apply (rule disjE) - apply (simp add: wn_triangle3 z) - apply (simp add: path_image_join winding_number_reversepath * wn_triangle3 z) + have "z \ closed_segment a c" "z \ closed_segment c b" + using assms by (meson convex_contains_segment convex_segment ends_in_segment subsetCE)+ + then show ?thesis + using assms + by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) +qed + +lemma winding_number_cong: + "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" + by (simp add: winding_number_def winding_number_prop_def pathstart_def pathfinish_def) + +lemma winding_number_constI: + assumes "c\z" "\t. \0\t; t\1\ \ g t = c" + shows "winding_number g z = 0" +proof - + have "winding_number g z = winding_number (linepath c c) z" + apply (rule winding_number_cong) + using assms unfolding linepath_def by auto + moreover have "winding_number (linepath c c) z =0" + apply (rule winding_number_trivial) + using assms by auto + ultimately show ?thesis by auto +qed + +lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" + unfolding winding_number_def +proof (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) + fix n e g + assume "0 < e" and g: "winding_number_prop p z e g n" + then show "\r. winding_number_prop (\w. p w - z) 0 e r n" + by (rule_tac x="\t. g t - z" in exI) + (force simp: winding_number_prop_def contour_integral_integral valid_path_def path_defs + vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) +next + fix n e g + assume "0 < e" and g: "winding_number_prop (\w. p w - z) 0 e g n" + then show "\r. winding_number_prop p z e r n" + apply (rule_tac x="\t. g t + z" in exI) + apply (simp add: winding_number_prop_def contour_integral_integral valid_path_def path_defs + piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) + apply (force simp: algebra_simps) done qed -subsection\Winding numbers for simple closed paths\ +subsubsection\<^marker>\tag unimportant\ \Some lemmas about negating a path\ -lemma winding_number_from_innerpath: - assumes "simple_path c1" and c1: "pathstart c1 = a" "pathfinish c1 = b" - and "simple_path c2" and c2: "pathstart c2 = a" "pathfinish c2 = b" - and "simple_path c" and c: "pathstart c = a" "pathfinish c = b" - and c1c2: "path_image c1 \ path_image c2 = {a,b}" - and c1c: "path_image c1 \ path_image c = {a,b}" - and c2c: "path_image c2 \ path_image c = {a,b}" - and ne_12: "path_image c \ inside(path_image c1 \ path_image c2) \ {}" - and z: "z \ inside(path_image c1 \ path_image c)" - and wn_d: "winding_number (c1 +++ reversepath c) z = d" - and "a \ b" "d \ 0" - obtains "z \ inside(path_image c1 \ path_image c2)" "winding_number (c1 +++ reversepath c2) z = d" -proof - - obtain 0: "inside(path_image c1 \ path_image c) \ inside(path_image c2 \ path_image c) = {}" - and 1: "inside(path_image c1 \ path_image c) \ inside(path_image c2 \ path_image c) \ - (path_image c - {a,b}) = inside(path_image c1 \ path_image c2)" - by (rule split_inside_simple_closed_curve - [OF \simple_path c1\ c1 \simple_path c2\ c2 \simple_path c\ c \a \ b\ c1c2 c1c c2c ne_12]) - have znot: "z \ path_image c" "z \ path_image c1" "z \ path_image c2" - using union_with_outside z 1 by auto - have wn_cc2: "winding_number (c +++ reversepath c2) z = 0" - apply (rule winding_number_zero_in_outside) - apply (simp_all add: \simple_path c2\ c c2 \simple_path c\ simple_path_imp_path path_image_join) - by (metis "0" ComplI UnE disjoint_iff_not_equal sup.commute union_with_inside z znot) - show ?thesis - proof - show "z \ inside (path_image c1 \ path_image c2)" - using "1" z by blast - have "winding_number c1 z - winding_number c z = d " - using assms znot - by (metis wn_d winding_number_join simple_path_imp_path winding_number_reversepath add.commute path_image_reversepath path_reversepath pathstart_reversepath uminus_add_conv_diff) - then show "winding_number (c1 +++ reversepath c2) z = d" - using wn_cc2 by (simp add: winding_number_join simple_path_imp_path assms znot winding_number_reversepath) - qed -qed +lemma valid_path_negatepath: "valid_path \ \ valid_path (uminus \ \)" + unfolding o_def using piecewise_C1_differentiable_neg valid_path_def by blast -lemma simple_closed_path_wn1: - fixes a::complex and e::real - assumes "0 < e" - and sp_pl: "simple_path(p +++ linepath (a - e) (a + e))" - and psp: "pathstart p = a + e" - and pfp: "pathfinish p = a - e" - and disj: "ball a e \ path_image p = {}" -obtains z where "z \ inside (path_image (p +++ linepath (a - e) (a + e)))" - "cmod (winding_number (p +++ linepath (a - e) (a + e)) z) = 1" +lemma has_contour_integral_negatepath: + assumes \: "valid_path \" and cint: "((\z. f (- z)) has_contour_integral - i) \" + shows "(f has_contour_integral i) (uminus \ \)" proof - - have "arc p" and arc_lp: "arc (linepath (a - e) (a + e))" - and pap: "path_image p \ path_image (linepath (a - e) (a + e)) \ {pathstart p, a-e}" - using simple_path_join_loop_eq [of "linepath (a - e) (a + e)" p] assms by auto - have mid_eq_a: "midpoint (a - e) (a + e) = a" - by (simp add: midpoint_def) - then have "a \ path_image(p +++ linepath (a - e) (a + e))" - apply (simp add: assms path_image_join) - by (metis midpoint_in_closed_segment) - have "a \ frontier(inside (path_image(p +++ linepath (a - e) (a + e))))" - apply (simp add: assms Jordan_inside_outside) - apply (simp_all add: assms path_image_join) - by (metis mid_eq_a midpoint_in_closed_segment) - with \0 < e\ obtain c where c: "c \ inside (path_image(p +++ linepath (a - e) (a + e)))" - and dac: "dist a c < e" - by (auto simp: frontier_straddle) - then have "c \ path_image(p +++ linepath (a - e) (a + e))" - using inside_no_overlap by blast - then have "c \ path_image p" - "c \ closed_segment (a - of_real e) (a + of_real e)" - by (simp_all add: assms path_image_join) - with \0 < e\ dac have "c \ affine hull {a - of_real e, a + of_real e}" - by (simp add: segment_as_ball not_le) - with \0 < e\ have *: "\ collinear {a - e, c,a + e}" - using collinear_3_affine_hull [of "a-e" "a+e"] by (auto simp: insert_commute) - have 13: "1/3 + 1/3 + 1/3 = (1::real)" by simp - have "(1/3) *\<^sub>R (a - of_real e) + (1/3) *\<^sub>R c + (1/3) *\<^sub>R (a + of_real e) \ interior(convex hull {a - e, c, a + e})" - using interior_convex_hull_3_minimal [OF * DIM_complex] - by clarsimp (metis 13 zero_less_divide_1_iff zero_less_numeral) - then obtain z where z: "z \ interior(convex hull {a - e, c, a + e})" by force - have [simp]: "z \ closed_segment (a - e) c" - by (metis DIM_complex Diff_iff IntD2 inf_sup_absorb interior_of_triangle z) - have [simp]: "z \ closed_segment (a + e) (a - e)" - by (metis DIM_complex DiffD2 Un_iff interior_of_triangle z) - have [simp]: "z \ closed_segment c (a + e)" - by (metis (no_types, lifting) DIM_complex DiffD2 Un_insert_right inf_sup_aci(5) insertCI interior_of_triangle mk_disjoint_insert z) - show thesis - proof - have "norm (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z) = 1" - using winding_number_triangle [OF z] by simp - have zin: "z \ inside (path_image (linepath (a + e) (a - e)) \ path_image p)" - and zeq: "winding_number (linepath (a + e) (a - e) +++ reversepath p) z = - winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" - proof (rule winding_number_from_innerpath - [of "linepath (a + e) (a - e)" "a+e" "a-e" p - "linepath (a + e) c +++ linepath c (a - e)" z - "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z"]) - show sp_aec: "simple_path (linepath (a + e) c +++ linepath c (a - e))" - proof (rule arc_imp_simple_path [OF arc_join]) - show "arc (linepath (a + e) c)" - by (metis \c \ path_image p\ arc_linepath pathstart_in_path_image psp) - show "arc (linepath c (a - e))" - by (metis \c \ path_image p\ arc_linepath pathfinish_in_path_image pfp) - show "path_image (linepath (a + e) c) \ path_image (linepath c (a - e)) \ {pathstart (linepath c (a - e))}" - by clarsimp (metis "*" IntI Int_closed_segment closed_segment_commute singleton_iff) - qed auto - show "simple_path p" - using \arc p\ arc_simple_path by blast - show sp_ae2: "simple_path (linepath (a + e) (a - e))" - using \arc p\ arc_distinct_ends pfp psp by fastforce - show pa: "pathfinish (linepath (a + e) (a - e)) = a - e" - "pathstart (linepath (a + e) c +++ linepath c (a - e)) = a + e" - "pathfinish (linepath (a + e) c +++ linepath c (a - e)) = a - e" - "pathstart p = a + e" "pathfinish p = a - e" - "pathstart (linepath (a + e) (a - e)) = a + e" - by (simp_all add: assms) - show 1: "path_image (linepath (a + e) (a - e)) \ path_image p = {a + e, a - e}" - proof - show "path_image (linepath (a + e) (a - e)) \ path_image p \ {a + e, a - e}" - using pap closed_segment_commute psp segment_convex_hull by fastforce - show "{a + e, a - e} \ path_image (linepath (a + e) (a - e)) \ path_image p" - using pap pathfinish_in_path_image pathstart_in_path_image pfp psp by fastforce - qed - show 2: "path_image (linepath (a + e) (a - e)) \ path_image (linepath (a + e) c +++ linepath c (a - e)) = - {a + e, a - e}" (is "?lhs = ?rhs") - proof - have "\ collinear {c, a + e, a - e}" - using * by (simp add: insert_commute) - then have "convex hull {a + e, a - e} \ convex hull {a + e, c} = {a + e}" - "convex hull {a + e, a - e} \ convex hull {c, a - e} = {a - e}" - by (metis (full_types) Int_closed_segment insert_commute segment_convex_hull)+ - then show "?lhs \ ?rhs" - by (metis Int_Un_distrib equalityD1 insert_is_Un path_image_join path_image_linepath path_join_eq path_linepath segment_convex_hull simple_path_def sp_aec) - show "?rhs \ ?lhs" - using segment_convex_hull by (simp add: path_image_join) - qed - have "path_image p \ path_image (linepath (a + e) c) \ {a + e}" - proof (clarsimp simp: path_image_join) - fix x - assume "x \ path_image p" and x_ac: "x \ closed_segment (a + e) c" - then have "dist x a \ e" - by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less) - with x_ac dac \e > 0\ show "x = a + e" - by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a]) - qed - moreover - have "path_image p \ path_image (linepath c (a - e)) \ {a - e}" - proof (clarsimp simp: path_image_join) - fix x - assume "x \ path_image p" and x_ac: "x \ closed_segment c (a - e)" - then have "dist x a \ e" - by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less) - with x_ac dac \e > 0\ show "x = a - e" - by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a]) - qed - ultimately - have "path_image p \ path_image (linepath (a + e) c +++ linepath c (a - e)) \ {a + e, a - e}" - by (force simp: path_image_join) - then show 3: "path_image p \ path_image (linepath (a + e) c +++ linepath c (a - e)) = {a + e, a - e}" - apply (rule equalityI) - apply (clarsimp simp: path_image_join) - apply (metis pathstart_in_path_image psp pathfinish_in_path_image pfp) - done - show 4: "path_image (linepath (a + e) c +++ linepath c (a - e)) \ - inside (path_image (linepath (a + e) (a - e)) \ path_image p) \ {}" - apply (clarsimp simp: path_image_join segment_convex_hull disjoint_iff_not_equal) - by (metis (no_types, hide_lams) UnI1 Un_commute c closed_segment_commute ends_in_segment(1) path_image_join - path_image_linepath pathstart_linepath pfp segment_convex_hull) - show zin_inside: "z \ inside (path_image (linepath (a + e) (a - e)) \ - path_image (linepath (a + e) c +++ linepath c (a - e)))" - apply (simp add: path_image_join) - by (metis z inside_of_triangle DIM_complex Un_commute closed_segment_commute) - show 5: "winding_number - (linepath (a + e) (a - e) +++ reversepath (linepath (a + e) c +++ linepath c (a - e))) z = - winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" - by (simp add: reversepath_joinpaths path_image_join winding_number_join) - show 6: "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z \ 0" - by (simp add: winding_number_triangle z) - show "winding_number (linepath (a + e) (a - e) +++ reversepath p) z = - winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" - by (metis 1 2 3 4 5 6 pa sp_aec sp_ae2 \arc p\ \simple_path p\ arc_distinct_ends winding_number_from_innerpath zin_inside) - qed (use assms \e > 0\ in auto) - show "z \ inside (path_image (p +++ linepath (a - e) (a + e)))" - using zin by (simp add: assms path_image_join Un_commute closed_segment_commute) - then have "cmod (winding_number (p +++ linepath (a - e) (a + e)) z) = - cmod ((winding_number(reversepath (p +++ linepath (a - e) (a + e))) z))" - apply (subst winding_number_reversepath) - using simple_path_imp_path sp_pl apply blast - apply (metis IntI emptyE inside_no_overlap) - by (simp add: inside_def) - also have "... = cmod (winding_number(linepath (a + e) (a - e) +++ reversepath p) z)" - by (simp add: pfp reversepath_joinpaths) - also have "... = cmod (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z)" - by (simp add: zeq) - also have "... = 1" - using z by (simp add: interior_of_triangle winding_number_triangle) - finally show "cmod (winding_number (p +++ linepath (a - e) (a + e)) z) = 1" . + obtain S where cont: "continuous_on {0..1} \" and "finite S" and diff: "\ C1_differentiable_on {0..1} - S" + using \ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + have "((\x. - (f (- \ x) * vector_derivative \ (at x within {0..1}))) has_integral i) {0..1}" + using cint by (auto simp: has_contour_integral_def dest: has_integral_neg) + then + have "((\x. f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1})) has_integral i) {0..1}" + proof (rule rev_iffD1 [OF _ has_integral_spike_eq]) + show "negligible S" + by (simp add: \finite S\ negligible_finite) + show "f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1}) = + - (f (- \ x) * vector_derivative \ (at x within {0..1}))" + if "x \ {0..1} - S" for x + proof - + have "vector_derivative (uminus \ \) (at x within cbox 0 1) = - vector_derivative \ (at x within cbox 0 1)" + proof (rule vector_derivative_within_cbox) + show "(uminus \ \ has_vector_derivative - vector_derivative \ (at x within cbox 0 1)) (at x within cbox 0 1)" + using that unfolding o_def + by (metis C1_differentiable_on_eq UNIV_I diff differentiable_subset has_vector_derivative_minus subsetI that vector_derivative_works) + qed (use that in auto) + then show ?thesis + by simp + qed qed -qed - -lemma simple_closed_path_wn2: - fixes a::complex and d e::real - assumes "0 < d" "0 < e" - and sp_pl: "simple_path(p +++ linepath (a - d) (a + e))" - and psp: "pathstart p = a + e" - and pfp: "pathfinish p = a - d" -obtains z where "z \ inside (path_image (p +++ linepath (a - d) (a + e)))" - "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1" -proof - - have [simp]: "a + of_real x \ closed_segment (a - \) (a - \) \ x \ closed_segment (-\) (-\)" for x \ \::real - using closed_segment_translation_eq [of a] - by (metis (no_types, hide_lams) add_uminus_conv_diff of_real_minus of_real_closed_segment) - have [simp]: "a - of_real x \ closed_segment (a + \) (a + \) \ -x \ closed_segment \ \" for x \ \::real - by (metis closed_segment_translation_eq diff_conv_add_uminus of_real_closed_segment of_real_minus) - have "arc p" and arc_lp: "arc (linepath (a - d) (a + e))" and "path p" - and pap: "path_image p \ closed_segment (a - d) (a + e) \ {a+e, a-d}" - using simple_path_join_loop_eq [of "linepath (a - d) (a + e)" p] assms arc_imp_path by auto - have "0 \ closed_segment (-d) e" - using \0 < d\ \0 < e\ closed_segment_eq_real_ivl by auto - then have "a \ path_image (linepath (a - d) (a + e))" - using of_real_closed_segment [THEN iffD2] - by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) - then have "a \ path_image p" - using \0 < d\ \0 < e\ pap by auto - then obtain k where "0 < k" and k: "ball a k \ (path_image p) = {}" - using \0 < e\ \path p\ not_on_path_ball by blast - define kde where "kde \ (min k (min d e)) / 2" - have "0 < kde" "kde < k" "kde < d" "kde < e" - using \0 < k\ \0 < d\ \0 < e\ by (auto simp: kde_def) - let ?q = "linepath (a + kde) (a + e) +++ p +++ linepath (a - d) (a - kde)" - have "- kde \ closed_segment (-d) e" - using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto - then have a_diff_kde: "a - kde \ closed_segment (a - d) (a + e)" - using of_real_closed_segment [THEN iffD2] - by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) - then have clsub2: "closed_segment (a - d) (a - kde) \ closed_segment (a - d) (a + e)" - by (simp add: subset_closed_segment) - then have "path_image p \ closed_segment (a - d) (a - kde) \ {a + e, a - d}" - using pap by force - moreover - have "a + e \ path_image p \ closed_segment (a - d) (a - kde)" - using \0 < kde\ \kde < d\ \0 < e\ by (auto simp: closed_segment_eq_real_ivl) - ultimately have sub_a_diff_d: "path_image p \ closed_segment (a - d) (a - kde) \ {a - d}" - by blast - have "kde \ closed_segment (-d) e" - using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto - then have a_diff_kde: "a + kde \ closed_segment (a - d) (a + e)" - using of_real_closed_segment [THEN iffD2] - by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment) - then have clsub1: "closed_segment (a + kde) (a + e) \ closed_segment (a - d) (a + e)" - by (simp add: subset_closed_segment) - then have "closed_segment (a + kde) (a + e) \ path_image p \ {a + e, a - d}" - using pap by force - moreover - have "closed_segment (a + kde) (a + e) \ closed_segment (a - d) (a - kde) = {}" - proof (clarsimp intro!: equals0I) - fix y - assume y1: "y \ closed_segment (a + kde) (a + e)" - and y2: "y \ closed_segment (a - d) (a - kde)" - obtain u where u: "y = a + of_real u" and "0 < u" - using y1 \0 < kde\ \kde < e\ \0 < e\ apply (clarsimp simp: in_segment) - apply (rule_tac u = "(1 - u)*kde + u*e" in that) - apply (auto simp: scaleR_conv_of_real algebra_simps) - by (meson le_less_trans less_add_same_cancel2 less_eq_real_def mult_left_mono) - moreover - obtain v where v: "y = a + of_real v" and "v \ 0" - using y2 \0 < kde\ \0 < d\ \0 < e\ apply (clarsimp simp: in_segment) - apply (rule_tac v = "- ((1 - u)*d + u*kde)" in that) - apply (force simp: scaleR_conv_of_real algebra_simps) - by (meson less_eq_real_def neg_le_0_iff_le segment_bound_lemma) - ultimately show False - by auto - qed - moreover have "a - d \ closed_segment (a + kde) (a + e)" - using \0 < kde\ \kde < d\ \0 < e\ by (auto simp: closed_segment_eq_real_ivl) - ultimately have sub_a_plus_e: - "closed_segment (a + kde) (a + e) \ (path_image p \ closed_segment (a - d) (a - kde)) - \ {a + e}" - by auto - have "kde \ closed_segment (-kde) e" - using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto - then have a_add_kde: "a + kde \ closed_segment (a - kde) (a + e)" - using of_real_closed_segment [THEN iffD2] - by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment) - have "closed_segment (a - kde) (a + kde) \ closed_segment (a + kde) (a + e) = {a + kde}" - by (metis a_add_kde Int_closed_segment) - moreover - have "path_image p \ closed_segment (a - kde) (a + kde) = {}" - proof (rule equals0I, clarify) - fix y assume "y \ path_image p" "y \ closed_segment (a - kde) (a + kde)" - with equals0D [OF k, of y] \0 < kde\ \kde < k\ show False - by (auto simp: dist_norm dest: dist_decreases_closed_segment [where c=a]) - qed - moreover - have "- kde \ closed_segment (-d) kde" - using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto - then have a_diff_kde': "a - kde \ closed_segment (a - d) (a + kde)" - using of_real_closed_segment [THEN iffD2] - by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) - then have "closed_segment (a - d) (a - kde) \ closed_segment (a - kde) (a + kde) = {a - kde}" - by (metis Int_closed_segment) - ultimately - have pa_subset_pm_kde: "path_image ?q \ closed_segment (a - kde) (a + kde) \ {a - kde, a + kde}" - by (auto simp: path_image_join assms) - have ge_kde1: "\y. x = a + y \ y \ kde" if "x \ closed_segment (a + kde) (a + e)" for x - using that \kde < e\ mult_le_cancel_left - apply (auto simp: in_segment) - apply (rule_tac x="(1-u)*kde + u*e" in exI) - apply (fastforce simp: algebra_simps scaleR_conv_of_real) - done - have ge_kde2: "\y. x = a + y \ y \ -kde" if "x \ closed_segment (a - d) (a - kde)" for x - using that \kde < d\ affine_ineq - apply (auto simp: in_segment) - apply (rule_tac x="- ((1-u)*d + u*kde)" in exI) - apply (fastforce simp: algebra_simps scaleR_conv_of_real) - done - have notin_paq: "x \ path_image ?q" if "dist a x < kde" for x - using that using \0 < kde\ \kde < d\ \kde < k\ - apply (auto simp: path_image_join assms dist_norm dest!: ge_kde1 ge_kde2) - by (meson k disjoint_iff_not_equal le_less_trans less_eq_real_def mem_ball that) - obtain z where zin: "z \ inside (path_image (?q +++ linepath (a - kde) (a + kde)))" - and z1: "cmod (winding_number (?q +++ linepath (a - kde) (a + kde)) z) = 1" - proof (rule simple_closed_path_wn1 [of kde ?q a]) - show "simple_path (?q +++ linepath (a - kde) (a + kde))" - proof (intro simple_path_join_loop conjI) - show "arc ?q" - proof (rule arc_join) - show "arc (linepath (a + kde) (a + e))" - using \kde < e\ \arc p\ by (force simp: pfp) - show "arc (p +++ linepath (a - d) (a - kde))" - using \kde < d\ \kde < e\ \arc p\ sub_a_diff_d by (force simp: pfp intro: arc_join) - qed (auto simp: psp pfp path_image_join sub_a_plus_e) - show "arc (linepath (a - kde) (a + kde))" - using \0 < kde\ by auto - qed (use pa_subset_pm_kde in auto) - qed (use \0 < kde\ notin_paq in auto) - have eq: "path_image (?q +++ linepath (a - kde) (a + kde)) = path_image (p +++ linepath (a - d) (a + e))" - (is "?lhs = ?rhs") - proof - show "?lhs \ ?rhs" - using clsub1 clsub2 apply (auto simp: path_image_join assms) - by (meson subsetCE subset_closed_segment) - show "?rhs \ ?lhs" - apply (simp add: path_image_join assms Un_ac) - by (metis Un_closed_segment Un_assoc a_diff_kde a_diff_kde' le_supI2 subset_refl) - qed - show thesis - proof - show zzin: "z \ inside (path_image (p +++ linepath (a - d) (a + e)))" - by (metis eq zin) - then have znotin: "z \ path_image p" - by (metis ComplD Un_iff inside_Un_outside path_image_join pathfinish_linepath pathstart_reversepath pfp reversepath_linepath) - have znotin_de: "z \ closed_segment (a - d) (a + kde)" - by (metis ComplD Un_iff Un_closed_segment a_diff_kde inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin) - have "winding_number (linepath (a - d) (a + e)) z = - winding_number (linepath (a - d) (a + kde)) z + winding_number (linepath (a + kde) (a + e)) z" - apply (rule winding_number_split_linepath) - apply (simp add: a_diff_kde) - by (metis ComplD Un_iff inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin) - also have "... = winding_number (linepath (a + kde) (a + e)) z + - (winding_number (linepath (a - d) (a - kde)) z + - winding_number (linepath (a - kde) (a + kde)) z)" - by (simp add: winding_number_split_linepath [of "a-kde", symmetric] znotin_de a_diff_kde') - finally have "winding_number (p +++ linepath (a - d) (a + e)) z = - winding_number p z + winding_number (linepath (a + kde) (a + e)) z + - (winding_number (linepath (a - d) (a - kde)) z + - winding_number (linepath (a - kde) (a + kde)) z)" - by (metis (no_types, lifting) ComplD Un_iff \arc p\ add.assoc arc_imp_path eq path_image_join path_join_path_ends path_linepath simple_path_imp_path sp_pl union_with_outside winding_number_join zin) - also have "... = winding_number ?q z + winding_number (linepath (a - kde) (a + kde)) z" - using \path p\ znotin assms zzin clsub1 - apply (subst winding_number_join, auto) - apply (metis (no_types, hide_lams) ComplD Un_iff contra_subsetD inside_Un_outside path_image_join path_image_linepath pathstart_linepath) - apply (metis Un_iff Un_closed_segment a_diff_kde' not_in_path_image_join path_image_linepath znotin_de) - by (metis Un_iff Un_closed_segment a_diff_kde' path_image_linepath path_linepath pathstart_linepath winding_number_join znotin_de) - also have "... = winding_number (?q +++ linepath (a - kde) (a + kde)) z" - using \path p\ assms zin - apply (subst winding_number_join [symmetric], auto) - apply (metis ComplD Un_iff path_image_join pathfinish_join pathfinish_linepath pathstart_linepath union_with_outside) - by (metis Un_iff Un_closed_segment a_diff_kde' znotin_de) - finally have "winding_number (p +++ linepath (a - d) (a + e)) z = - winding_number (?q +++ linepath (a - kde) (a + kde)) z" . - then show "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1" - by (simp add: z1) - qed + then show ?thesis by (simp add: has_contour_integral_def) qed -lemma simple_closed_path_wn3: - fixes p :: "real \ complex" - assumes "simple_path p" and loop: "pathfinish p = pathstart p" - obtains z where "z \ inside (path_image p)" "cmod (winding_number p z) = 1" +lemma winding_number_negatepath: + assumes \: "valid_path \" and 0: "0 \ path_image \" + shows "winding_number(uminus \ \) 0 = winding_number \ 0" proof - - have ins: "inside(path_image p) \ {}" "open(inside(path_image p))" - "connected(inside(path_image p))" - and out: "outside(path_image p) \ {}" "open(outside(path_image p))" - "connected(outside(path_image p))" - and bo: "bounded(inside(path_image p))" "\ bounded(outside(path_image p))" - and ins_out: "inside(path_image p) \ outside(path_image p) = {}" - "inside(path_image p) \ outside(path_image p) = - path_image p" - and fro: "frontier(inside(path_image p)) = path_image p" - "frontier(outside(path_image p)) = path_image p" - using Jordan_inside_outside [OF assms] by auto - obtain a where a: "a \ inside(path_image p)" - using \inside (path_image p) \ {}\ by blast - obtain d::real where "0 < d" and d_fro: "a - d \ frontier (inside (path_image p))" - and d_int: "\\. \0 \ \; \ < d\ \ (a - \) \ inside (path_image p)" - apply (rule ray_to_frontier [of "inside (path_image p)" a "-1"]) - using \bounded (inside (path_image p))\ \open (inside (path_image p))\ a interior_eq - apply (auto simp: of_real_def) - done - obtain e::real where "0 < e" and e_fro: "a + e \ frontier (inside (path_image p))" - and e_int: "\\. \0 \ \; \ < e\ \ (a + \) \ inside (path_image p)" - apply (rule ray_to_frontier [of "inside (path_image p)" a 1]) - using \bounded (inside (path_image p))\ \open (inside (path_image p))\ a interior_eq - apply (auto simp: of_real_def) - done - obtain t0 where "0 \ t0" "t0 \ 1" and pt: "p t0 = a - d" - using a d_fro fro by (auto simp: path_image_def) - obtain q where "simple_path q" and q_ends: "pathstart q = a - d" "pathfinish q = a - d" - and q_eq_p: "path_image q = path_image p" - and wn_q_eq_wn_p: "\z. z \ inside(path_image p) \ winding_number q z = winding_number p z" - proof - show "simple_path (shiftpath t0 p)" - by (simp add: pathstart_shiftpath pathfinish_shiftpath - simple_path_shiftpath path_image_shiftpath \0 \ t0\ \t0 \ 1\ assms) - show "pathstart (shiftpath t0 p) = a - d" - using pt by (simp add: \t0 \ 1\ pathstart_shiftpath) - show "pathfinish (shiftpath t0 p) = a - d" - by (simp add: \0 \ t0\ loop pathfinish_shiftpath pt) - show "path_image (shiftpath t0 p) = path_image p" - by (simp add: \0 \ t0\ \t0 \ 1\ loop path_image_shiftpath) - show "winding_number (shiftpath t0 p) z = winding_number p z" - if "z \ inside (path_image p)" for z - by (metis ComplD Un_iff \0 \ t0\ \t0 \ 1\ \simple_path p\ atLeastAtMost_iff inside_Un_outside - loop simple_path_imp_path that winding_number_shiftpath) - qed - have ad_not_ae: "a - d \ a + e" - by (metis \0 < d\ \0 < e\ add.left_inverse add_left_cancel add_uminus_conv_diff - le_add_same_cancel2 less_eq_real_def not_less of_real_add of_real_def of_real_eq_0_iff pt) - have ad_ae_q: "{a - d, a + e} \ path_image q" - using \path_image q = path_image p\ d_fro e_fro fro(1) by auto - have ada: "open_segment (a - d) a \ inside (path_image p)" - proof (clarsimp simp: in_segment) - fix u::real assume "0 < u" "u < 1" - with d_int have "a - (1 - u) * d \ inside (path_image p)" - by (metis \0 < d\ add.commute diff_add_cancel left_diff_distrib' less_add_same_cancel2 less_eq_real_def mult.left_neutral zero_less_mult_iff) - then show "(1 - u) *\<^sub>R (a - d) + u *\<^sub>R a \ inside (path_image p)" - by (simp add: diff_add_eq of_real_def real_vector.scale_right_diff_distrib) - qed - have aae: "open_segment a (a + e) \ inside (path_image p)" - proof (clarsimp simp: in_segment) - fix u::real assume "0 < u" "u < 1" - with e_int have "a + u * e \ inside (path_image p)" - by (meson \0 < e\ less_eq_real_def mult_less_cancel_right2 not_less zero_less_mult_iff) - then show "(1 - u) *\<^sub>R a + u *\<^sub>R (a + e) \ inside (path_image p)" - apply (simp add: algebra_simps) - by (simp add: diff_add_eq of_real_def real_vector.scale_right_diff_distrib) - qed - have "complex_of_real (d * d + (e * e + d * (e + e))) \ 0" - using ad_not_ae - by (metis \0 < d\ \0 < e\ add_strict_left_mono less_add_same_cancel1 not_sum_squares_lt_zero - of_real_eq_0_iff zero_less_double_add_iff_zero_less_single_add zero_less_mult_iff) - then have a_in_de: "a \ open_segment (a - d) (a + e)" - using ad_not_ae \0 < d\ \0 < e\ - apply (auto simp: in_segment algebra_simps scaleR_conv_of_real) - apply (rule_tac x="d / (d+e)" in exI) - apply (auto simp: field_simps) + have "(/) 1 contour_integrable_on \" + using "0" \ contour_integrable_inversediff by fastforce + then have "((\z. 1/z) has_contour_integral contour_integral \ ((/) 1)) \" + by (rule has_contour_integral_integral) + then have "((\z. 1 / - z) has_contour_integral - contour_integral \ ((/) 1)) \" + using has_contour_integral_neg by auto + then show ?thesis + using assms + apply (simp add: winding_number_valid_path valid_path_negatepath image_def path_defs) + apply (simp add: contour_integral_unique has_contour_integral_negatepath) done - then have "open_segment (a - d) (a + e) \ inside (path_image p)" - using ada a aae Un_open_segment [of a "a-d" "a+e"] by blast - then have "path_image q \ open_segment (a - d) (a + e) = {}" - using inside_no_overlap by (fastforce simp: q_eq_p) - with ad_ae_q have paq_Int_cs: "path_image q \ closed_segment (a - d) (a + e) = {a - d, a + e}" - by (simp add: closed_segment_eq_open) - obtain t where "0 \ t" "t \ 1" and qt: "q t = a + e" - using a e_fro fro ad_ae_q by (auto simp: path_defs) - then have "t \ 0" - by (metis ad_not_ae pathstart_def q_ends(1)) - then have "t \ 1" - by (metis ad_not_ae pathfinish_def q_ends(2) qt) - have q01: "q 0 = a - d" "q 1 = a - d" - using q_ends by (auto simp: pathstart_def pathfinish_def) - obtain z where zin: "z \ inside (path_image (subpath t 0 q +++ linepath (a - d) (a + e)))" - and z1: "cmod (winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z) = 1" - proof (rule simple_closed_path_wn2 [of d e "subpath t 0 q" a], simp_all add: q01) - show "simple_path (subpath t 0 q +++ linepath (a - d) (a + e))" - proof (rule simple_path_join_loop, simp_all add: qt q01) - have "inj_on q (closed_segment t 0)" - using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ \t \ 1\ - by (fastforce simp: simple_path_def inj_on_def closed_segment_eq_real_ivl) - then show "arc (subpath t 0 q)" - using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ - by (simp add: arc_subpath_eq simple_path_imp_path) - show "arc (linepath (a - d) (a + e))" - by (simp add: ad_not_ae) - show "path_image (subpath t 0 q) \ closed_segment (a - d) (a + e) \ {a + e, a - d}" - using qt paq_Int_cs \simple_path q\ \0 \ t\ \t \ 1\ - by (force simp: dest: rev_subsetD [OF _ path_image_subpath_subset] intro: simple_path_imp_path) - qed - qed (auto simp: \0 < d\ \0 < e\ qt) - have pa01_Un: "path_image (subpath 0 t q) \ path_image (subpath 1 t q) = path_image q" - unfolding path_image_subpath - using \0 \ t\ \t \ 1\ by (force simp: path_image_def image_iff) - with paq_Int_cs have pa_01q: - "(path_image (subpath 0 t q) \ path_image (subpath 1 t q)) \ closed_segment (a - d) (a + e) = {a - d, a + e}" - by metis - have z_notin_ed: "z \ closed_segment (a + e) (a - d)" - using zin q01 by (simp add: path_image_join closed_segment_commute inside_def) - have z_notin_0t: "z \ path_image (subpath 0 t q)" - by (metis (no_types, hide_lams) IntI Un_upper1 subsetD empty_iff inside_no_overlap path_image_join - path_image_subpath_commute pathfinish_subpath pathstart_def pathstart_linepath q_ends(1) qt subpath_trivial zin) - have [simp]: "- winding_number (subpath t 0 q) z = winding_number (subpath 0 t q) z" - by (metis \0 \ t\ \simple_path q\ \t \ 1\ atLeastAtMost_iff zero_le_one - path_image_subpath_commute path_subpath real_eq_0_iff_le_ge_0 - reversepath_subpath simple_path_imp_path winding_number_reversepath z_notin_0t) - obtain z_in_q: "z \ inside(path_image q)" - and wn_q: "winding_number (subpath 0 t q +++ subpath t 1 q) z = - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z" - proof (rule winding_number_from_innerpath - [of "subpath 0 t q" "a-d" "a+e" "subpath 1 t q" "linepath (a - d) (a + e)" - z "- winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z"], - simp_all add: q01 qt pa01_Un reversepath_subpath) - show "simple_path (subpath 0 t q)" "simple_path (subpath 1 t q)" - by (simp_all add: \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ \t \ 1\ simple_path_subpath) - show "simple_path (linepath (a - d) (a + e))" - using ad_not_ae by blast - show "path_image (subpath 0 t q) \ path_image (subpath 1 t q) = {a - d, a + e}" (is "?lhs = ?rhs") - proof - show "?lhs \ ?rhs" - using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 1\ q_ends qt q01 - by (force simp: pathfinish_def qt simple_path_def path_image_subpath) - show "?rhs \ ?lhs" - using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) - qed - show "path_image (subpath 0 t q) \ closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs") - proof - show "?lhs \ ?rhs" using paq_Int_cs pa01_Un by fastforce - show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) - qed - show "path_image (subpath 1 t q) \ closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs") - proof - show "?lhs \ ?rhs" by (auto simp: pa_01q [symmetric]) - show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) - qed - show "closed_segment (a - d) (a + e) \ inside (path_image q) \ {}" - using a a_in_de open_closed_segment pa01_Un q_eq_p by fastforce - show "z \ inside (path_image (subpath 0 t q) \ closed_segment (a - d) (a + e))" - by (metis path_image_join path_image_linepath path_image_subpath_commute pathfinish_subpath pathstart_linepath q01(1) zin) - show "winding_number (subpath 0 t q +++ linepath (a + e) (a - d)) z = - - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z" - using z_notin_ed z_notin_0t \0 \ t\ \simple_path q\ \t \ 1\ - by (simp add: simple_path_imp_path qt q01 path_image_subpath_commute closed_segment_commute winding_number_join winding_number_reversepath [symmetric]) - show "- d \ e" - using ad_not_ae by auto - show "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z \ 0" - using z1 by auto - qed - show ?thesis - proof - show "z \ inside (path_image p)" - using q_eq_p z_in_q by auto - then have [simp]: "z \ path_image q" - by (metis disjoint_iff_not_equal inside_no_overlap q_eq_p) - have [simp]: "z \ path_image (subpath 1 t q)" - using inside_def pa01_Un z_in_q by fastforce - have "winding_number(subpath 0 t q +++ subpath t 1 q) z = winding_number(subpath 0 1 q) z" - using z_notin_0t \0 \ t\ \simple_path q\ \t \ 1\ - by (simp add: simple_path_imp_path qt path_image_subpath_commute winding_number_join winding_number_subpath_combine) - with wn_q have "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z = - winding_number q z" - by auto - with z1 have "cmod (winding_number q z) = 1" - by simp - with z1 wn_q_eq_wn_p show "cmod (winding_number p z) = 1" - using z1 wn_q_eq_wn_p by (simp add: \z \ inside (path_image p)\) - qed qed -proposition simple_closed_path_winding_number_inside: - assumes "simple_path \" - obtains "\z. z \ inside(path_image \) \ winding_number \ z = 1" - | "\z. z \ inside(path_image \) \ winding_number \ z = -1" -proof (cases "pathfinish \ = pathstart \") - case True - have "path \" - by (simp add: assms simple_path_imp_path) - then have const: "winding_number \ constant_on inside(path_image \)" - proof (rule winding_number_constant) - show "connected (inside(path_image \))" - by (simp add: Jordan_inside_outside True assms) - qed (use inside_no_overlap True in auto) - obtain z where zin: "z \ inside (path_image \)" and z1: "cmod (winding_number \ z) = 1" - using simple_closed_path_wn3 [of \] True assms by blast - have "winding_number \ z \ \" - using zin integer_winding_number [OF \path \\ True] inside_def by blast - with z1 consider "winding_number \ z = 1" | "winding_number \ z = -1" - apply (auto simp: Ints_def abs_if split: if_split_asm) - by (metis of_int_1 of_int_eq_iff of_int_minus) - with that const zin show ?thesis - unfolding constant_on_def by metis -next - case False - then show ?thesis - using inside_simple_curve_imp_closed assms that(2) by blast -qed - -lemma simple_closed_path_abs_winding_number_inside: - assumes "simple_path \" "z \ inside(path_image \)" - shows "\Re (winding_number \ z)\ = 1" - by (metis assms norm_minus_cancel norm_one one_complex.simps(1) real_norm_def simple_closed_path_winding_number_inside uminus_complex.simps(1)) - -lemma simple_closed_path_norm_winding_number_inside: - assumes "simple_path \" "z \ inside(path_image \)" - shows "norm (winding_number \ z) = 1" -proof - - have "pathfinish \ = pathstart \" - using assms inside_simple_curve_imp_closed by blast - with assms integer_winding_number have "winding_number \ z \ \" - by (simp add: inside_def simple_path_def) - then show ?thesis - by (metis assms norm_minus_cancel norm_one simple_closed_path_winding_number_inside) -qed +lemma contour_integrable_negatepath: + assumes \: "valid_path \" and pi: "(\z. f (- z)) contour_integrable_on \" + shows "f contour_integrable_on (uminus \ \)" + by (metis \ add.inverse_inverse contour_integrable_on_def has_contour_integral_negatepath pi) -lemma simple_closed_path_winding_number_cases: - "\simple_path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ {-1,0,1}" -apply (simp add: inside_Un_outside [of "path_image \", symmetric, unfolded set_eq_iff Set.Compl_iff] del: inside_Un_outside) - apply (rule simple_closed_path_winding_number_inside) - using simple_path_def winding_number_zero_in_outside by blast+ - -lemma simple_closed_path_winding_number_pos: - "\simple_path \; pathfinish \ = pathstart \; z \ path_image \; 0 < Re(winding_number \ z)\ - \ winding_number \ z = 1" -using simple_closed_path_winding_number_cases - by fastforce - -subsection \Winding number for rectangular paths\ - -definition\<^marker>\tag important\ rectpath where - "rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3) - in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)" - -lemma path_rectpath [simp, intro]: "path (rectpath a b)" - by (simp add: Let_def rectpath_def) - -lemma valid_path_rectpath [simp, intro]: "valid_path (rectpath a b)" - by (simp add: Let_def rectpath_def) +(* A combined theorem deducing several things piecewise.*) +lemma winding_number_join_pos_combined: + "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); + valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ + \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" + by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) -lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1" - by (simp add: rectpath_def Let_def) - -lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1" - by (simp add: rectpath_def Let_def) - -lemma simple_path_rectpath [simp, intro]: - assumes "Re a1 \ Re a3" "Im a1 \ Im a3" - shows "simple_path (rectpath a1 a3)" - unfolding rectpath_def Let_def using assms - by (intro simple_path_join_loop arc_join arc_linepath) - (auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im) -lemma path_image_rectpath: - assumes "Re a1 \ Re a3" "Im a1 \ Im a3" - shows "path_image (rectpath a1 a3) = - {z. Re z \ {Re a1, Re a3} \ Im z \ {Im a1..Im a3}} \ - {z. Im z \ {Im a1, Im a3} \ Re z \ {Re a1..Re a3}}" (is "?lhs = ?rhs") + +subsubsection\<^marker>\tag unimportant\ \Useful sufficient conditions for the winding number to be positive\ + +lemma Re_winding_number: + "\valid_path \; z \ path_image \\ + \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" +by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) + +lemma winding_number_pos_le: + assumes \: "valid_path \" "z \ path_image \" + and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" + shows "0 \ Re(winding_number \ z)" proof - - define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" - have "?lhs = closed_segment a1 a2 \ closed_segment a2 a3 \ - closed_segment a4 a3 \ closed_segment a1 a4" - by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute - a2_def a4_def Un_assoc) - also have "\ = ?rhs" using assms - by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def - closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl) - finally show ?thesis . + have ge0: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x + using ge by (simp add: Complex.Im_divide algebra_simps x) + let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" + let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" + have hi: "(?vd has_integral ?int z) (cbox 0 1)" + unfolding box_real + apply (subst has_contour_integral [symmetric]) + using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) + have "0 \ Im (?int z)" + proof (rule has_integral_component_nonneg [of \, simplified]) + show "\x. x \ cbox 0 1 \ 0 \ Im (if 0 < x \ x < 1 then ?vd x else 0)" + by (force simp: ge0) + show "((\x. if 0 < x \ x < 1 then ?vd x else 0) has_integral ?int z) (cbox 0 1)" + by (rule has_integral_spike_interior [OF hi]) simp + qed + then show ?thesis + by (simp add: Re_winding_number [OF \] field_simps) qed -lemma path_image_rectpath_subset_cbox: - assumes "Re a \ Re b" "Im a \ Im b" - shows "path_image (rectpath a b) \ cbox a b" - using assms by (auto simp: path_image_rectpath in_cbox_complex_iff) - -lemma path_image_rectpath_inter_box: - assumes "Re a \ Re b" "Im a \ Im b" - shows "path_image (rectpath a b) \ box a b = {}" - using assms by (auto simp: path_image_rectpath in_box_complex_iff) - -lemma path_image_rectpath_cbox_minus_box: - assumes "Re a \ Re b" "Im a \ Im b" - shows "path_image (rectpath a b) = cbox a b - box a b" - using assms by (auto simp: path_image_rectpath in_cbox_complex_iff - in_box_complex_iff) - -proposition winding_number_rectpath: - assumes "z \ box a1 a3" - shows "winding_number (rectpath a1 a3) z = 1" +lemma winding_number_pos_lt_lemma: + assumes \: "valid_path \" "z \ path_image \" + and e: "0 < e" + and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" + shows "0 < Re(winding_number \ z)" proof - - from assms have less: "Re a1 < Re a3" "Im a1 < Im a3" - by (auto simp: in_box_complex_iff) - define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" - let ?l1 = "linepath a1 a2" and ?l2 = "linepath a2 a3" - and ?l3 = "linepath a3 a4" and ?l4 = "linepath a4 a1" - from assms and less have "z \ path_image (rectpath a1 a3)" - by (auto simp: path_image_rectpath_cbox_minus_box) - also have "path_image (rectpath a1 a3) = - path_image ?l1 \ path_image ?l2 \ path_image ?l3 \ path_image ?l4" - by (simp add: rectpath_def Let_def path_image_join Un_assoc a2_def a4_def) - finally have "z \ \" . - moreover have "\l\{?l1,?l2,?l3,?l4}. Re (winding_number l z) > 0" - unfolding ball_simps HOL.simp_thms a2_def a4_def - by (intro conjI; (rule winding_number_linepath_pos_lt; - (insert assms, auto simp: a2_def a4_def in_box_complex_iff mult_neg_neg))+) - ultimately have "Re (winding_number (rectpath a1 a3) z) > 0" - by (simp add: winding_number_join path_image_join rectpath_def Let_def a2_def a4_def) - thus "winding_number (rectpath a1 a3) z = 1" using assms less - by (intro simple_closed_path_winding_number_pos simple_path_rectpath) - (auto simp: path_image_rectpath_cbox_minus_box) + let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" + let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" + have hi: "(?vd has_integral ?int z) (cbox 0 1)" + unfolding box_real + apply (subst has_contour_integral [symmetric]) + using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) + have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" + proof (rule has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified]) + show "((\x. if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e) has_integral ?int z) {0..1}" + by (rule has_integral_spike_interior [OF hi, simplified box_real]) (use e in simp) + show "\x. 0 \ x \ x \ 1 \ + e \ Im (if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e)" + by (simp add: ge) + qed (use has_integral_const_real [of _ 0 1] in auto) + with e show ?thesis + by (simp add: Re_winding_number [OF \] field_simps) qed -proposition winding_number_rectpath_outside: - assumes "Re a1 \ Re a3" "Im a1 \ Im a3" - assumes "z \ cbox a1 a3" - shows "winding_number (rectpath a1 a3) z = 0" - using assms by (intro winding_number_zero_outside[OF _ _ _ assms(3)] - path_image_rectpath_subset_cbox) simp_all - -text\A per-function version for continuous logs, a kind of monodromy\ -proposition\<^marker>\tag unimportant\ winding_number_compose_exp: - assumes "path p" - shows "winding_number (exp \ p) 0 = (pathfinish p - pathstart p) / (2 * of_real pi * \)" +lemma winding_number_pos_lt: + assumes \: "valid_path \" "z \ path_image \" + and e: "0 < e" + and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" + shows "0 < Re (winding_number \ z)" proof - - obtain e where "0 < e" and e: "\t. t \ {0..1} \ e \ norm(exp(p t))" - proof - have "closed (path_image (exp \ p))" - by (simp add: assms closed_path_image holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image) - then show "0 < setdist {0} (path_image (exp \ p))" - by (metis exp_not_eq_zero imageE image_comp infdist_eq_setdist infdist_pos_not_in_closed path_defs(4) path_image_nonempty) - next - fix t::real - assume "t \ {0..1}" - have "setdist {0} (path_image (exp \ p)) \ dist 0 (exp (p t))" - apply (rule setdist_le_dist) - using \t \ {0..1}\ path_image_def by fastforce+ - then show "setdist {0} (path_image (exp \ p)) \ cmod (exp (p t))" - by simp - qed - have "bounded (path_image p)" - by (simp add: assms bounded_path_image) - then obtain B where "0 < B" and B: "path_image p \ cball 0 B" - by (meson bounded_pos mem_cball_0 subsetI) - let ?B = "cball (0::complex) (B+1)" - have "uniformly_continuous_on ?B exp" - using holomorphic_on_exp holomorphic_on_imp_continuous_on - by (force intro: compact_uniformly_continuous) - then obtain d where "d > 0" - and d: "\x x'. \x\?B; x'\?B; dist x' x < d\ \ norm (exp x' - exp x) < e" - using \e > 0\ by (auto simp: uniformly_continuous_on_def dist_norm) - then have "min 1 d > 0" + have bm: "bounded ((\w. w - z) ` (path_image \))" + using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) + then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" + using bounded_pos [THEN iffD1, OF bm] by blast + { fix x::real assume x: "0 < x" "x < 1" + then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] + by (simp add: path_image_def power2_eq_square mult_mono') + with x have "\ x \ z" using \ + using path_image_def by fastforce + then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" + using B ge [OF x] B2 e + apply (rule_tac y="e / (cmod (\ x - z))\<^sup>2" in order_trans) + apply (auto simp: divide_left_mono divide_right_mono) + done + then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" + by (simp add: complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) + } note * = this + show ?thesis + using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) +qed + +subsection\The winding number is an integer\ + +text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, + Also on page 134 of Serge Lang's book with the name title, etc.\ + +lemma exp_fg: + fixes z::complex + assumes g: "(g has_vector_derivative g') (at x within s)" + and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" + and z: "g x \ z" + shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" +proof - + have *: "(exp \ (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" + using assms unfolding has_vector_derivative_def scaleR_conv_of_real + by (auto intro!: derivative_eq_intros) + show ?thesis + apply (rule has_vector_derivative_eq_rhs) + using z + apply (auto intro!: derivative_eq_intros * [unfolded o_def] g) + done +qed + +lemma winding_number_exp_integral: + fixes z::complex + assumes \: "\ piecewise_C1_differentiable_on {a..b}" + and ab: "a \ b" + and z: "z \ \ ` {a..b}" + shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" + (is "?thesis1") + "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" + (is "?thesis2") +proof - + let ?D\ = "\x. vector_derivative \ (at x)" + have [simp]: "\x. a \ x \ x \ b \ \ x \ z" + using z by force + have cong: "continuous_on {a..b} \" + using \ by (simp add: piecewise_C1_differentiable_on_def) + obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" + using \ by (force simp: piecewise_C1_differentiable_on_def) + have \: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) + moreover have "{a<.. {a..b} - k" by force - then obtain g where pfg: "polynomial_function g" and "g 0 = p 0" "g 1 = p 1" - and gless: "\t. t \ {0..1} \ norm(g t - p t) < min 1 d" - using path_approx_polynomial_function [OF \path p\] \d > 0\ \0 < e\ - unfolding pathfinish_def pathstart_def by meson - have "winding_number (exp \ p) 0 = winding_number (exp \ g) 0" - proof (rule winding_number_nearby_paths_eq [symmetric]) - show "path (exp \ p)" "path (exp \ g)" - by (simp_all add: pfg assms holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image path_polynomial_function) - next - fix t :: "real" - assume t: "t \ {0..1}" - with gless have "norm(g t - p t) < 1" - using min_less_iff_conj by blast - moreover have ptB: "norm (p t) \ B" - using B t by (force simp: path_image_def) - ultimately have "cmod (g t) \ B + 1" - by (meson add_mono_thms_linordered_field(4) le_less_trans less_imp_le norm_triangle_sub) - with ptB gless t have "cmod ((exp \ g) t - (exp \ p) t) < e" - by (auto simp: dist_norm d) - with e t show "cmod ((exp \ g) t - (exp \ p) t) < cmod ((exp \ p) t - 0)" - by fastforce - qed (use \g 0 = p 0\ \g 1 = p 1\ in \auto simp: pathfinish_def pathstart_def\) - also have "... = 1 / (of_real (2 * pi) * \) * contour_integral (exp \ g) (\w. 1 / (w - 0))" - proof (rule winding_number_valid_path) - have "continuous_on (path_image g) (deriv exp)" - by (metis DERIV_exp DERIV_imp_deriv continuous_on_cong holomorphic_on_exp holomorphic_on_imp_continuous_on) - then show "valid_path (exp \ g)" - by (simp add: field_differentiable_within_exp pfg valid_path_compose valid_path_polynomial_function) - show "0 \ path_image (exp \ g)" - by (auto simp: path_image_def) + ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" + by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def at_within_open) + { fix w + assume "w \ z" + have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" + by (auto simp: dist_norm intro!: continuous_intros) + moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" + by (auto simp: intro!: derivative_eq_intros) + ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" + using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] + by (force simp: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) + } + then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" + by meson + have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" + unfolding integrable_on_def [symmetric] + proof (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \]]) + show "\d h. 0 < d \ + (\y. cmod (y - w) < d \ (h has_field_derivative inverse (y - z))(at y within - {z}))" + if "w \ - {z}" for w + apply (rule_tac x="norm(w - z)" in exI) + using that inverse_eq_divide has_field_derivative_at_within h + by (metis Compl_insert DiffD2 insertCI right_minus_eq zero_less_norm_iff) + qed simp + have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" + unfolding box_real [symmetric] divide_inverse_commute + by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) + with ab show ?thesis1 + by (simp add: divide_inverse_commute integral_def integrable_on_def) + { fix t + assume t: "t \ {a..b}" + have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" + using z by (auto intro!: continuous_intros simp: dist_norm) + have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" + unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) + obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ + (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" + using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] + by simp (auto simp: ball_def dist_norm that) + { fix x D + assume x: "x \ k" "a < x" "x < b" + then have "x \ interior ({a..b} - k)" + using open_subset_interior [OF \] by fastforce + then have con: "isCont ?D\ x" + using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) + then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" + by (rule continuous_at_imp_continuous_within) + have gdx: "\ differentiable at x" + using x by (simp add: g_diff_at) + have "\d. \x \ k; a < x; x < b; + (\ has_vector_derivative d) (at x); a \ t; t \ b\ + \ ((\x. integral {a..x} + (\x. ?D\ x / + (\ x - z))) has_vector_derivative + d / (\ x - z)) + (at x within {a..b})" + apply (rule has_vector_derivative_eq_rhs) + apply (rule integral_has_vector_derivative_continuous_at [where S = "{}", simplified]) + apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+ + done + then have "((\c. exp (- integral {a..c} (\x. ?D\ x / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) + (at x within {a..b})" + using x gdx t + apply (clarsimp simp add: differentiable_iff_scaleR) + apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_withinI) + apply (simp_all add: has_vector_derivative_def [symmetric]) + done + } note * = this + have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" + apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) + using t + apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous_1 [OF vg_int] simp add: ab)+ + done + } + with ab show ?thesis2 + by (simp add: divide_inverse_commute integral_def) +qed + +lemma winding_number_exp_2pi: + "\path p; z \ path_image p\ + \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" +using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def winding_number_prop_def + by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) + +lemma integer_winding_number_eq: + assumes \: "path \" and z: "z \ path_image \" + shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" +proof - + obtain p where p: "valid_path p" "z \ path_image p" + "pathstart p = pathstart \" "pathfinish p = pathfinish \" + and eq: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" + using winding_number [OF assms, of 1] unfolding winding_number_prop_def by auto + then have wneq: "winding_number \ z = winding_number p z" + using eq winding_number_valid_path by force + have iff: "(winding_number \ z \ \) \ (exp (contour_integral p (\w. 1 / (w - z))) = 1)" + using eq by (simp add: exp_eq_1 complex_is_Int_iff) + have "exp (contour_integral p (\w. 1 / (w - z))) = (\ 1 - z) / (\ 0 - z)" + using p winding_number_exp_integral(2) [of p 0 1 z] + apply (simp add: valid_path_def path_defs contour_integral_integral exp_minus field_split_simps) + by (metis path_image_def pathstart_def pathstart_in_path_image) + then have "winding_number p z \ \ \ pathfinish p = pathstart p" + using p wneq iff by (auto simp: path_defs) + then show ?thesis using p eq + by (auto simp: winding_number_valid_path) +qed + +theorem integer_winding_number: + "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" +by (metis integer_winding_number_eq) + + +text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) + We can thus bound the winding number of a path that doesn't intersect a given ray. \ + +lemma winding_number_pos_meets: + fixes z::complex + assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" + and w: "w \ z" + shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" +proof - + have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" + using z by (auto simp: path_image_def) + have [simp]: "z \ \ ` {0..1}" + using path_image_def z by auto + have gpd: "\ piecewise_C1_differentiable_on {0..1}" + using \ valid_path_def by blast + define r where "r = (w - z) / (\ 0 - z)" + have [simp]: "r \ 0" + using w z by (auto simp: r_def) + have cont: "continuous_on {0..1} + (\x. Im (integral {0..x} (\x. vector_derivative \ (at x) / (\ x - z))))" + by (intro continuous_intros indefinite_integral_continuous_1 winding_number_exp_integral [OF gpd]; simp) + have "Arg2pi r \ 2*pi" + by (simp add: Arg2pi less_eq_real_def) + also have "\ \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" + using 1 + apply (simp add: winding_number_valid_path [OF \ z] contour_integral_integral) + apply (simp add: Complex.Re_divide field_simps power2_eq_square) + done + finally have "Arg2pi r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . + then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" + by (simp add: Arg2pi_ge_0 cont IVT') + then obtain t where t: "t \ {0..1}" + and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" + by blast + define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" + have iArg: "Arg2pi r = Im i" + using eqArg by (simp add: i_def) + have gpdt: "\ piecewise_C1_differentiable_on {0..t}" + by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) + have "exp (- i) * (\ t - z) = \ 0 - z" + unfolding i_def + apply (rule winding_number_exp_integral [OF gpdt]) + using t z unfolding path_image_def by force+ + then have *: "\ t - z = exp i * (\ 0 - z)" + by (simp add: exp_minus field_simps) + then have "(w - z) = r * (\ 0 - z)" + by (simp add: r_def) + then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" + apply simp + apply (subst Complex_Transcendental.Arg2pi_eq [of r]) + apply (simp add: iArg) + using * apply (simp add: exp_eq_polar field_simps) + done + with t show ?thesis + by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) +qed + +lemma winding_number_big_meets: + fixes z::complex + assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" + and w: "w \ z" + shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" +proof - + { assume "Re (winding_number \ z) \ - 1" + then have "Re (winding_number (reversepath \) z) \ 1" + by (simp add: \ valid_path_imp_path winding_number_reversepath z) + moreover have "valid_path (reversepath \)" + using \ valid_path_imp_reverse by auto + moreover have "z \ path_image (reversepath \)" + by (simp add: z) + ultimately have "\a::real. 0 < a \ z + a*(w - z) \ path_image (reversepath \)" + using winding_number_pos_meets w by blast + then have ?thesis + by simp + } + then show ?thesis + using assms + by (simp add: abs_if winding_number_pos_meets split: if_split_asm) +qed + +lemma winding_number_less_1: + fixes z::complex + shows + "\valid_path \; z \ path_image \; w \ z; + \a::real. 0 < a \ z + a*(w - z) \ path_image \\ + \ Re(winding_number \ z) < 1" + by (auto simp: not_less dest: winding_number_big_meets) + +text\One way of proving that WN=1 for a loop.\ +lemma winding_number_eq_1: + fixes z::complex + assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" + and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" + shows "winding_number \ z = 1" +proof - + have "winding_number \ z \ Ints" + by (simp add: \ integer_winding_number loop valid_path_imp_path z) + then show ?thesis + using 0 2 by (auto simp: Ints_def) +qed + +subsection\Continuity of winding number and invariance on connected sets\ + +lemma continuous_at_winding_number: + fixes z::complex + assumes \: "path \" and z: "z \ path_image \" + shows "continuous (at z) (winding_number \)" +proof - + obtain e where "e>0" and cbg: "cball z e \ - path_image \" + using open_contains_cball [of "- path_image \"] z + by (force simp: closed_def [symmetric] closed_path_image [OF \]) + then have ppag: "path_image \ \ - cball z (e/2)" + by (force simp: cball_def dist_norm) + have oc: "open (- cball z (e / 2))" + by (simp add: closed_def [symmetric]) + obtain d where "d>0" and pi_eq: + "\h1 h2. \valid_path h1; valid_path h2; + (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); + pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ + \ + path_image h1 \ - cball z (e / 2) \ + path_image h2 \ - cball z (e / 2) \ + (\f. f holomorphic_on - cball z (e / 2) \ contour_integral h2 f = contour_integral h1 f)" + using contour_integral_nearby_ends [OF oc \ ppag] by metis + obtain p where p: "valid_path p" "z \ path_image p" + "pathstart p = pathstart \ \ pathfinish p = pathfinish \" + and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e / 2" + and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" + using winding_number [OF \ z, of "min d e / 2"] \d>0\ \e>0\ by (auto simp: winding_number_prop_def) + { fix w + assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" + then have wnotp: "w \ path_image p" + using cbg \d>0\ \e>0\ + apply (simp add: path_image_def cball_def dist_norm, clarify) + apply (frule pg) + apply (drule_tac c="\ x" in subsetD) + apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) + done + have wnotg: "w \ path_image \" + using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) + { fix k::real + assume k: "k>0" + then obtain q where q: "valid_path q" "w \ path_image q" + "pathstart q = pathstart \ \ pathfinish q = pathfinish \" + and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" + and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" + using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k + by (force simp: min_divide_distrib_right winding_number_prop_def) + have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" + apply (rule pi_eq [OF \valid_path q\ \valid_path p\, THEN conjunct2, THEN conjunct2, rule_format]) + apply (frule pg) + apply (frule qg) + using p q \d>0\ e2 + apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) + done + then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" + by (simp add: pi qi) + } note pip = this + have "path p" + using p by (simp add: valid_path_imp_path) + then have "winding_number p w = winding_number \ w" + apply (rule winding_number_unique [OF _ wnotp]) + apply (rule_tac x=p in exI) + apply (simp add: p wnotp min_divide_distrib_right pip winding_number_prop_def) + done + } note wnwn = this + obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" + using p open_contains_cball [of "- path_image p"] + by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) + obtain L + where "L>0" + and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); + \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ + cmod (contour_integral p f) \ L * B" + using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by force + { fix e::real and w::complex + assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" + then have [simp]: "w \ path_image p" + using cbp p(2) \0 < pe\ + by (force simp: dist_norm norm_minus_commute path_image_def cball_def) + have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = + contour_integral p (\x. 1/(x - w) - 1/(x - z))" + by (simp add: p contour_integrable_inversediff contour_integral_diff) + { fix x + assume pe: "3/4 * pe < cmod (z - x)" + have "cmod (w - x) < pe/4 + cmod (z - x)" + by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) + then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp + have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" + using norm_diff_triangle_le by blast + also have "\ < pe/4 + cmod (w - x)" + using w by (simp add: norm_minus_commute) + finally have "pe/2 < cmod (w - x)" + using pe by auto + then have "(pe/2)^2 < cmod (w - x) ^ 2" + apply (rule power_strict_mono) + using \pe>0\ by auto + then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" + by (simp add: power_divide) + have "8 * L * cmod (w - z) < e * pe\<^sup>2" + using w \L>0\ by (simp add: field_simps) + also have "\ < e * 4 * cmod (w - x) * cmod (w - x)" + using pe2 \e>0\ by (simp add: power2_eq_square) + also have "\ < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" + using wx + apply (rule mult_strict_left_mono) + using pe2 e not_less_iff_gr_or_eq by fastforce + finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" + by simp + also have "\ \ e * cmod (w - x) * cmod (z - x)" + using e by simp + finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . + have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" + apply (cases "x=z \ x=w") + using pe \pe>0\ w \L>0\ + apply (force simp: norm_minus_commute) + using wx w(2) \L>0\ pe pe2 Lwz + apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) + done + } note L_cmod_le = this + have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" + apply (rule L) + using \pe>0\ w + apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) + using \pe>0\ w \L>0\ + apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) + done + have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" + apply simp + apply (rule le_less_trans [OF *]) + using \L>0\ e + apply (force simp: field_simps) + done + then have "cmod (winding_number p w - winding_number p z) < e" + using pi_ge_two e + by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) + } note cmod_wn_diff = this + then have "isCont (winding_number p) z" + apply (simp add: continuous_at_eps_delta, clarify) + apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) + using \pe>0\ \L>0\ + apply (simp add: dist_norm cmod_wn_diff) + done + then show ?thesis + apply (rule continuous_transform_within [where d = "min d e / 2"]) + apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) + done +qed + +corollary continuous_on_winding_number: + "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" + by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) + +subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ + +lemma winding_number_constant: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" + shows "winding_number \ constant_on S" +proof - + have *: "1 \ cmod (winding_number \ y - winding_number \ z)" + if ne: "winding_number \ y \ winding_number \ z" and "y \ S" "z \ S" for y z + proof - + have "winding_number \ y \ \" "winding_number \ z \ \" + using that integer_winding_number [OF \ loop] sg \y \ S\ by auto + with ne show ?thesis + by (auto simp: Ints_def simp flip: of_int_diff) qed - also have "... = 1 / (of_real (2 * pi) * \) * integral {0..1} (\x. vector_derivative g (at x))" - proof (simp add: contour_integral_integral, rule integral_cong) - fix t :: "real" - assume t: "t \ {0..1}" - show "vector_derivative (exp \ g) (at t) / exp (g t) = vector_derivative g (at t)" - proof - - have "(exp \ g has_vector_derivative vector_derivative (exp \ g) (at t)) (at t)" - by (meson DERIV_exp differentiable_def field_vector_diff_chain_at has_vector_derivative_def - has_vector_derivative_polynomial_function pfg vector_derivative_works) - moreover have "(exp \ g has_vector_derivative vector_derivative g (at t) * exp (g t)) (at t)" - apply (rule field_vector_diff_chain_at) - apply (metis has_vector_derivative_polynomial_function pfg vector_derivative_at) - using DERIV_exp has_field_derivative_def apply blast + have cont: "continuous_on S (\w. winding_number \ w)" + using continuous_on_winding_number [OF \] sg + by (meson continuous_on_subset disjoint_eq_subset_Compl) + show ?thesis + using "*" zero_less_one + by (blast intro: continuous_discrete_range_constant [OF cs cont]) +qed + +lemma winding_number_eq: + "\path \; pathfinish \ = pathstart \; w \ S; z \ S; connected S; S \ path_image \ = {}\ + \ winding_number \ w = winding_number \ z" + using winding_number_constant by (metis constant_on_def) + +lemma open_winding_number_levelsets: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + shows "open {z. z \ path_image \ \ winding_number \ z = k}" +proof - + have opn: "open (- path_image \)" + by (simp add: closed_path_image \ open_Compl) + { fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" + obtain e where e: "e>0" "ball z e \ - path_image \" + using open_contains_ball [of "- path_image \"] opn z + by blast + have "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" + apply (rule_tac x=e in exI) + using e apply (simp add: dist_norm ball_def norm_minus_commute) + apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where S = "ball z e"]) + done + } then + show ?thesis + by (auto simp: open_dist) +qed + +proposition winding_number_zero_in_outside: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" + shows "winding_number \ z = 0" +proof - + obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" + using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto + obtain w::complex where w: "w \ ball 0 (B + 1)" + by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) + have "- ball 0 (B + 1) \ outside (path_image \)" + apply (rule outside_subset_convex) + using B subset_ball by auto + then have wout: "w \ outside (path_image \)" + using w by blast + moreover have "winding_number \ constant_on outside (path_image \)" + using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside + by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) + ultimately have "winding_number \ z = winding_number \ w" + by (metis (no_types, hide_lams) constant_on_def z) + also have "\ = 0" + proof - + have wnot: "w \ path_image \" using wout by (simp add: outside_def) + { fix e::real assume "0" "pathfinish p = pathfinish \" + and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" + and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" + using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force + have pip: "path_image p \ ball 0 (B + 1)" + using B + apply (clarsimp simp add: path_image_def dist_norm ball_def) + apply (frule (1) pg1) + apply (fastforce dest: norm_add_less) done - ultimately show ?thesis - by (simp add: divide_simps, rule vector_derivative_unique_at) - qed - qed - also have "... = (pathfinish p - pathstart p) / (2 * of_real pi * \)" - proof - - have "((\x. vector_derivative g (at x)) has_integral g 1 - g 0) {0..1}" - apply (rule fundamental_theorem_of_calculus [OF zero_le_one]) - by (metis has_vector_derivative_at_within has_vector_derivative_polynomial_function pfg vector_derivative_at) + then have "w \ path_image p" using w by blast + then have "\p. valid_path p \ w \ path_image p \ + pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ + (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" + apply (rule_tac x=p in exI) + apply (simp add: p valid_path_polynomial_function) + apply (intro conjI) + using pge apply (simp add: norm_minus_commute) + apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) + apply (rule holomorphic_intros | simp add: dist_norm)+ + using mem_ball_0 w apply blast + using p apply (simp_all add: valid_path_polynomial_function loop pip) + done + } then show ?thesis - apply (simp add: pathfinish_def pathstart_def) - using \g 0 = p 0\ \g 1 = p 1\ by auto + by (auto intro: winding_number_unique [OF \] simp add: winding_number_prop_def wnot) qed finally show ?thesis . qed -subsection\<^marker>\tag unimportant\ \The winding number defines a continuous logarithm for the path itself\ +corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" + by (rule winding_number_zero_in_outside) + (auto simp: pathfinish_def pathstart_def path_polynomial_function) -lemma winding_number_as_continuous_log: - assumes "path p" and \: "\ \ path_image p" - obtains q where "path q" - "pathfinish q - pathstart q = 2 * of_real pi * \ * winding_number p \" - "\t. t \ {0..1} \ p t = \ + exp(q t)" +corollary\<^marker>\tag unimportant\ winding_number_zero_outside: + "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" + by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) + +lemma winding_number_zero_at_infinity: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + shows "\B. \z. B \ norm z \ winding_number \ z = 0" proof - - let ?q = "\t. 2 * of_real pi * \ * winding_number(subpath 0 t p) \ + Ln(pathstart p - \)" - show ?thesis - proof - have *: "continuous (at t within {0..1}) (\x. winding_number (subpath 0 x p) \)" - if t: "t \ {0..1}" for t - proof - - let ?B = "ball (p t) (norm(p t - \))" - have "p t \ \" - using path_image_def that \ by blast - then have "simply_connected ?B" - by (simp add: convex_imp_simply_connected) - then have "\f::complex\complex. continuous_on ?B f \ (\\ \ ?B. f \ \ 0) - \ (\g. continuous_on ?B g \ (\\ \ ?B. f \ = exp (g \)))" - by (simp add: simply_connected_eq_continuous_log) - moreover have "continuous_on ?B (\w. w - \)" - by (intro continuous_intros) - moreover have "(\z \ ?B. z - \ \ 0)" - by (auto simp: dist_norm) - ultimately obtain g where contg: "continuous_on ?B g" - and geq: "\z. z \ ?B \ z - \ = exp (g z)" by blast - obtain d where "0 < d" and d: - "\x. \x \ {0..1}; dist x t < d\ \ dist (p x) (p t) < cmod (p t - \)" - using \path p\ t unfolding path_def continuous_on_iff - by (metis \p t \ \\ right_minus_eq zero_less_norm_iff) - have "((\x. winding_number (\w. subpath 0 x p w - \) 0 - - winding_number (\w. subpath 0 t p w - \) 0) \ 0) - (at t within {0..1})" - proof (rule Lim_transform_within [OF _ \d > 0\]) - have "continuous (at t within {0..1}) (g o p)" - proof (rule continuous_within_compose) - show "continuous (at t within {0..1}) p" - using \path p\ continuous_on_eq_continuous_within path_def that by blast - show "continuous (at (p t) within p ` {0..1}) g" - by (metis (no_types, lifting) open_ball UNIV_I \p t \ \\ centre_in_ball contg continuous_on_eq_continuous_at continuous_within_topological right_minus_eq zero_less_norm_iff) - qed - with LIM_zero have "((\u. (g (subpath t u p 1) - g (subpath t u p 0))) \ 0) (at t within {0..1})" - by (auto simp: subpath_def continuous_within o_def) - then show "((\u. (g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \)) \ 0) - (at t within {0..1})" - by (simp add: tendsto_divide_zero) - show "(g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \) = - winding_number (\w. subpath 0 u p w - \) 0 - winding_number (\w. subpath 0 t p w - \) 0" - if "u \ {0..1}" "0 < dist u t" "dist u t < d" for u - proof - - have "closed_segment t u \ {0..1}" - using closed_segment_eq_real_ivl t that by auto - then have piB: "path_image(subpath t u p) \ ?B" - apply (clarsimp simp add: path_image_subpath_gen) - by (metis subsetD le_less_trans \dist u t < d\ d dist_commute dist_in_closed_segment) - have *: "path (g \ subpath t u p)" - apply (rule path_continuous_image) - using \path p\ t that apply auto[1] - using piB contg continuous_on_subset by blast - have "(g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \) - = winding_number (exp \ g \ subpath t u p) 0" - using winding_number_compose_exp [OF *] - by (simp add: pathfinish_def pathstart_def o_assoc) - also have "... = winding_number (\w. subpath t u p w - \) 0" - proof (rule winding_number_cong) - have "exp(g y) = y - \" if "y \ (subpath t u p) ` {0..1}" for y - by (metis that geq path_image_def piB subset_eq) - then show "\x. \0 \ x; x \ 1\ \ (exp \ g \ subpath t u p) x = subpath t u p x - \" - by auto - qed - also have "... = winding_number (\w. subpath 0 u p w - \) 0 - - winding_number (\w. subpath 0 t p w - \) 0" - apply (simp add: winding_number_offset [symmetric]) - using winding_number_subpath_combine [OF \path p\ \, of 0 t u] \t \ {0..1}\ \u \ {0..1}\ - by (simp add: add.commute eq_diff_eq) - finally show ?thesis . - qed - qed - then show ?thesis - by (subst winding_number_offset) (simp add: continuous_within LIM_zero_iff) - qed - show "path ?q" - unfolding path_def - by (intro continuous_intros) (simp add: continuous_on_eq_continuous_within *) - - have "\ \ p 0" - by (metis \ pathstart_def pathstart_in_path_image) - then show "pathfinish ?q - pathstart ?q = 2 * of_real pi * \ * winding_number p \" - by (simp add: pathfinish_def pathstart_def) - show "p t = \ + exp (?q t)" if "t \ {0..1}" for t - proof - - have "path (subpath 0 t p)" - using \path p\ that by auto - moreover - have "\ \ path_image (subpath 0 t p)" - using \ [unfolded path_image_def] that by (auto simp: path_image_subpath) - ultimately show ?thesis - using winding_number_exp_2pi [of "subpath 0 t p" \] \\ \ p 0\ - by (auto simp: exp_add algebra_simps pathfinish_def pathstart_def subpath_def) - qed - qed + obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" + using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto + then show ?thesis + apply (rule_tac x="B+1" in exI, clarify) + apply (rule winding_number_zero_outside [OF \ convex_cball [of 0 B] loop]) + apply (meson less_add_one mem_cball_0 not_le order_trans) + using ball_subset_cball by blast qed -subsection \Winding number equality is the same as path/loop homotopy in C - {0}\ +lemma winding_number_zero_point: + "\path \; convex s; pathfinish \ = pathstart \; open s; path_image \ \ s\ + \ \z. z \ s \ winding_number \ z = 0" + using outside_compact_in_open [of "path_image \" s] path_image_nonempty winding_number_zero_in_outside + by (fastforce simp add: compact_path_image) -lemma winding_number_homotopic_loops_null_eq: - assumes "path p" and \: "\ \ path_image p" - shows "winding_number p \ = 0 \ (\a. homotopic_loops (-{\}) p (\t. a))" - (is "?lhs = ?rhs") -proof - assume [simp]: ?lhs - obtain q where "path q" - and qeq: "pathfinish q - pathstart q = 2 * of_real pi * \ * winding_number p \" - and peq: "\t. t \ {0..1} \ p t = \ + exp(q t)" - using winding_number_as_continuous_log [OF assms] by blast - have *: "homotopic_with_canon (\r. pathfinish r = pathstart r) - {0..1} (-{\}) ((\w. \ + exp w) \ q) ((\w. \ + exp w) \ (\t. 0))" - proof (rule homotopic_with_compose_continuous_left) - show "homotopic_with_canon (\f. pathfinish ((\w. \ + exp w) \ f) = pathstart ((\w. \ + exp w) \ f)) - {0..1} UNIV q (\t. 0)" - proof (rule homotopic_with_mono, simp_all add: pathfinish_def pathstart_def) - have "homotopic_loops UNIV q (\t. 0)" - by (rule homotopic_loops_linear) (use qeq \path q\ in \auto simp: path_defs\) - then have "homotopic_with (\r. r 1 = r 0) (top_of_set {0..1}) euclidean q (\t. 0)" - by (simp add: homotopic_loops_def pathfinish_def pathstart_def) - then show "homotopic_with (\h. exp (h 1) = exp (h 0)) (top_of_set {0..1}) euclidean q (\t. 0)" - by (rule homotopic_with_mono) simp - qed - show "continuous_on UNIV (\w. \ + exp w)" - by (rule continuous_intros)+ - show "range (\w. \ + exp w) \ -{\}" - by auto + +text\If a path winds round a set, it winds rounds its inside.\ +lemma winding_number_around_inside: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + and cls: "closed s" and cos: "connected s" and s_disj: "s \ path_image \ = {}" + and z: "z \ s" and wn_nz: "winding_number \ z \ 0" and w: "w \ s \ inside s" + shows "winding_number \ w = winding_number \ z" +proof - + have ssb: "s \ inside(path_image \)" + proof + fix x :: complex + assume "x \ s" + hence "x \ path_image \" + by (meson disjoint_iff_not_equal s_disj) + thus "x \ inside (path_image \)" + using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) +qed + show ?thesis + apply (rule winding_number_eq [OF \ loop w]) + using z apply blast + apply (simp add: cls connected_with_inside cos) + apply (simp add: Int_Un_distrib2 s_disj, safe) + by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \ inside_no_overlap) + qed + +subsection \The real part of winding numbers\ + +text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ +lemma winding_number_subpath_continuous: + assumes \: "valid_path \" and z: "z \ path_image \" + shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" +proof - + have *: "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = + winding_number (subpath 0 x \) z" + if x: "0 \ x" "x \ 1" for x + proof - + have "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = + 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" + using assms x + apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) + done + also have "\ = winding_number (subpath 0 x \) z" + apply (subst winding_number_valid_path) + using assms x + apply (simp_all add: path_image_subpath valid_path_subpath) + by (force simp: path_image_def) + finally show ?thesis . qed - then have "homotopic_with_canon (\r. pathfinish r = pathstart r) {0..1} (-{\}) p (\x. \ + 1)" - by (rule homotopic_with_eq) (auto simp: o_def peq pathfinish_def pathstart_def) - then have "homotopic_loops (-{\}) p (\t. \ + 1)" - by (simp add: homotopic_loops_def) - then show ?rhs .. -next - assume ?rhs - then obtain a where "homotopic_loops (-{\}) p (\t. a)" .. - then have "winding_number p \ = winding_number (\t. a) \" "a \ \" - using winding_number_homotopic_loops homotopic_loops_imp_subset by (force simp:)+ - moreover have "winding_number (\t. a) \ = 0" - by (metis winding_number_zero_const \a \ \\) - ultimately show ?lhs by metis -qed - -lemma winding_number_homotopic_paths_null_explicit_eq: - assumes "path p" and \: "\ \ path_image p" - shows "winding_number p \ = 0 \ homotopic_paths (-{\}) p (linepath (pathstart p) (pathstart p))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - apply (auto simp: winding_number_homotopic_loops_null_eq [OF assms]) - apply (rule homotopic_loops_imp_homotopic_paths_null) - apply (simp add: linepath_refl) + show ?thesis + apply (rule continuous_on_eq + [where f = "\x. 1 / (2*pi*\) * + integral {0..x} (\t. 1/(\ t - z) * vector_derivative \ (at t))"]) + apply (rule continuous_intros)+ + apply (rule indefinite_integral_continuous_1) + apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) + using assms + apply (simp add: *) done -next - assume ?rhs - then show ?lhs - by (metis \ pathstart_in_path_image winding_number_homotopic_paths winding_number_trivial) qed -lemma winding_number_homotopic_paths_null_eq: - assumes "path p" and \: "\ \ path_image p" - shows "winding_number p \ = 0 \ (\a. homotopic_paths (-{\}) p (\t. a))" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - by (auto simp: winding_number_homotopic_paths_null_explicit_eq [OF assms] linepath_refl) -next - assume ?rhs - then show ?lhs - by (metis \ homotopic_paths_imp_pathfinish pathfinish_def pathfinish_in_path_image winding_number_homotopic_paths winding_number_zero_const) +lemma winding_number_ivt_pos: + assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" + shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" + apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) + apply (rule winding_number_subpath_continuous [OF \ z]) + using assms + apply (auto simp: path_image_def image_def) + done + +lemma winding_number_ivt_neg: + assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" + shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" + apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) + apply (rule winding_number_subpath_continuous [OF \ z]) + using assms + apply (auto simp: path_image_def image_def) + done + +lemma winding_number_ivt_abs: + assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" + shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" + using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] + by force + +lemma winding_number_lt_half_lemma: + assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" + shows "Re(winding_number \ z) < 1/2" +proof - + { assume "Re(winding_number \ z) \ 1/2" + then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" + using winding_number_ivt_pos [OF \ z, of "1/2"] by auto + have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" + using winding_number_exp_2pi [of "subpath 0 t \" z] + apply (simp add: t \ valid_path_imp_path) + using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) + have "b < a \ \ 0" + proof - + have "\ 0 \ {c. b < a \ c}" + by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) + thus ?thesis + by blast + qed + moreover have "b < a \ \ t" + proof - + have "\ t \ {c. b < a \ c}" + by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) + thus ?thesis + by blast + qed + ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az + by (simp add: inner_diff_right)+ + then have False + by (simp add: gt inner_mult_right mult_less_0_iff) + } + then show ?thesis by force qed -proposition winding_number_homotopic_paths_eq: - assumes "path p" and \p: "\ \ path_image p" - and "path q" and \q: "\ \ path_image q" - and qp: "pathstart q = pathstart p" "pathfinish q = pathfinish p" - shows "winding_number p \ = winding_number q \ \ homotopic_paths (-{\}) p q" - (is "?lhs = ?rhs") -proof - assume ?lhs - then have "winding_number (p +++ reversepath q) \ = 0" - using assms by (simp add: winding_number_join winding_number_reversepath) - moreover - have "path (p +++ reversepath q)" "\ \ path_image (p +++ reversepath q)" - using assms by (auto simp: not_in_path_image_join) - ultimately obtain a where "homotopic_paths (- {\}) (p +++ reversepath q) (linepath a a)" - using winding_number_homotopic_paths_null_explicit_eq by blast - then show ?rhs - using homotopic_paths_imp_pathstart assms - by (fastforce simp add: dest: homotopic_paths_imp_homotopic_loops homotopic_paths_loop_parts) -next - assume ?rhs - then show ?lhs - by (simp add: winding_number_homotopic_paths) +lemma winding_number_lt_half: + assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" + shows "\Re (winding_number \ z)\ < 1/2" +proof - + have "z \ path_image \" using assms by auto + with assms show ?thesis + apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) + apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] + winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) + done qed -lemma winding_number_homotopic_loops_eq: - assumes "path p" and \p: "\ \ path_image p" - and "path q" and \q: "\ \ path_image q" - and loops: "pathfinish p = pathstart p" "pathfinish q = pathstart q" - shows "winding_number p \ = winding_number q \ \ homotopic_loops (-{\}) p q" - (is "?lhs = ?rhs") -proof - assume L: ?lhs - have "pathstart p \ \" "pathstart q \ \" - using \p \q by blast+ - moreover have "path_connected (-{\})" - by (simp add: path_connected_punctured_universe) - ultimately obtain r where "path r" and rim: "path_image r \ -{\}" - and pas: "pathstart r = pathstart p" and paf: "pathfinish r = pathstart q" - by (auto simp: path_connected_def) - then have "pathstart r \ \" by blast - have "homotopic_loops (- {\}) p (r +++ q +++ reversepath r)" - proof (rule homotopic_paths_imp_homotopic_loops) - show "homotopic_paths (- {\}) p (r +++ q +++ reversepath r)" - by (metis (mono_tags, hide_lams) \path r\ L \p \q \path p\ \path q\ homotopic_loops_conjugate loops not_in_path_image_join paf pas path_image_reversepath path_imp_reversepath path_join_eq pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath rim subset_Compl_singleton winding_number_homotopic_loops winding_number_homotopic_paths_eq) - qed (use loops pas in auto) - moreover have "homotopic_loops (- {\}) (r +++ q +++ reversepath r) q" - using rim \q by (auto simp: homotopic_loops_conjugate paf \path q\ \path r\ loops) - ultimately show ?rhs - using homotopic_loops_trans by metis -next - assume ?rhs - then show ?lhs - by (simp add: winding_number_homotopic_loops) +lemma winding_number_le_half: + assumes \: "valid_path \" and z: "z \ path_image \" + and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" + shows "\Re (winding_number \ z)\ \ 1/2" +proof - + { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" + have "isCont (winding_number \) z" + by (metis continuous_at_winding_number valid_path_imp_path \ z) + then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" + using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast + define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" + have *: "a \ z' \ b - d / 3 * cmod a" + unfolding z'_def inner_mult_right' divide_inverse + apply (simp add: field_split_simps algebra_simps dot_square_norm power2_eq_square anz) + apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) + done + have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" + using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) + then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" + by simp + then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" + using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp + then have wnz_12': "\Re (winding_number \ z')\ > 1/2" + by linarith + moreover have "\Re (winding_number \ z')\ < 1/2" + apply (rule winding_number_lt_half [OF \ *]) + using azb \d>0\ pag + apply (auto simp: add_strict_increasing anz field_split_simps dest!: subsetD) + done + ultimately have False + by simp + } + then show ?thesis by force qed -end +lemma winding_number_lt_half_linepath: "z \ closed_segment a b \ \Re (winding_number (linepath a b) z)\ < 1/2" + using separating_hyperplane_closed_point [of "closed_segment a b" z] + apply auto + apply (simp add: closed_segment_def) + apply (drule less_imp_le) + apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) + apply (auto simp: segment) + done + +text\ Positivity of WN for a linepath.\ +lemma winding_number_linepath_pos_lt: + assumes "0 < Im ((b - a) * cnj (b - z))" + shows "0 < Re(winding_number(linepath a b) z)" +proof - + have z: "z \ path_image (linepath a b)" + using assms + by (simp add: closed_segment_def) (force simp: algebra_simps) + show ?thesis + apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) + apply (simp add: linepath_def algebra_simps) + done +qed + +proposition winding_number_part_circlepath_pos_less: + assumes "s < t" and no: "norm(w - z) < r" + shows "0 < Re (winding_number(part_circlepath z r s t) w)" +proof - + have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) + note valid_path_part_circlepath + moreover have " w \ path_image (part_circlepath z r s t)" + using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) + moreover have "0 < r * (t - s) * (r - cmod (w - z))" + using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) + ultimately show ?thesis + apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) + apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) + apply (rule mult_left_mono)+ + using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] + apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) + using assms \0 < r\ by auto +qed + +subsection \Invariance of winding numbers under homotopy\ + +text\including the fact that it's +-1 inside a simple closed curve.\ + +lemma winding_number_homotopic_paths: + assumes "homotopic_paths (-{z}) g h" + shows "winding_number g z = winding_number h z" +proof - + have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto + moreover have pag: "z \ path_image g" and pah: "z \ path_image h" + using homotopic_paths_imp_subset [OF assms] by auto + ultimately obtain d e where "d > 0" "e > 0" + and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ + \ homotopic_paths (-{z}) g p" + and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ + \ homotopic_paths (-{z}) h q" + using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force + obtain p where p: + "valid_path p" "z \ path_image p" + "pathstart p = pathstart g" "pathfinish p = pathfinish g" + and gp_less:"\t\{0..1}. cmod (g t - p t) < d" + and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" + using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast + obtain q where q: + "valid_path q" "z \ path_image q" + "pathstart q = pathstart h" "pathfinish q = pathfinish h" + and hq_less: "\t\{0..1}. cmod (h t - q t) < e" + and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" + using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast + have "homotopic_paths (- {z}) g p" + by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) + moreover have "homotopic_paths (- {z}) h q" + by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) + ultimately have "homotopic_paths (- {z}) p q" + by (blast intro: homotopic_paths_trans homotopic_paths_sym assms) + then have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" + by (rule Cauchy_theorem_homotopic_paths) (auto intro!: holomorphic_intros simp: p q) + then show ?thesis + by (simp add: pap paq) +qed + +lemma winding_number_homotopic_loops: + assumes "homotopic_loops (-{z}) g h" + shows "winding_number g z = winding_number h z" +proof - + have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto + moreover have pag: "z \ path_image g" and pah: "z \ path_image h" + using homotopic_loops_imp_subset [OF assms] by auto + moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" + using homotopic_loops_imp_loop [OF assms] by auto + ultimately obtain d e where "d > 0" "e > 0" + and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ + \ homotopic_loops (-{z}) g p" + and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ + \ homotopic_loops (-{z}) h q" + using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force + obtain p where p: + "valid_path p" "z \ path_image p" + "pathstart p = pathstart g" "pathfinish p = pathfinish g" + and gp_less:"\t\{0..1}. cmod (g t - p t) < d" + and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" + using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast + obtain q where q: + "valid_path q" "z \ path_image q" + "pathstart q = pathstart h" "pathfinish q = pathfinish h" + and hq_less: "\t\{0..1}. cmod (h t - q t) < e" + and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" + using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast + have gp: "homotopic_loops (- {z}) g p" + by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) + have hq: "homotopic_loops (- {z}) h q" + by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) + have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" + proof (rule Cauchy_theorem_homotopic_loops) + show "homotopic_loops (- {z}) p q" + by (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) + qed (auto intro!: holomorphic_intros simp: p q) + then show ?thesis + by (simp add: pap paq) +qed + +lemma winding_number_paths_linear_eq: + "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; + \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ + \ winding_number h z = winding_number g z" + by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths) + +lemma winding_number_loops_linear_eq: + "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; + \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ + \ winding_number h z = winding_number g z" + by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops) + +lemma winding_number_nearby_paths_eq: + "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; + \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ + \ winding_number h z = winding_number g z" + by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) + +lemma winding_number_nearby_loops_eq: + "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; + \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ + \ winding_number h z = winding_number g z" + by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) + + +lemma winding_number_subpath_combine: + "\path g; z \ path_image g; + u \ {0..1}; v \ {0..1}; w \ {0..1}\ + \ winding_number (subpath u v g) z + winding_number (subpath v w g) z = + winding_number (subpath u w g) z" +apply (rule trans [OF winding_number_join [THEN sym] + winding_number_homotopic_paths [OF homotopic_join_subpaths]]) + using path_image_subpath_subset by auto + +subsection \Winding numbers of some simple paths\ + +lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" + apply (rule winding_number_unique_loop) + apply (simp_all add: sphere_def valid_path_imp_path) + apply (rule_tac x="circlepath z r" in exI) + apply (simp add: sphere_def contour_integral_circlepath) + done + +proposition winding_number_circlepath: + assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" +proof (cases "w = z") + case True then show ?thesis + using assms winding_number_circlepath_centre by auto +next + case False + have [simp]: "r > 0" + using assms le_less_trans norm_ge_zero by blast + define r' where "r' = norm(w - z)" + have "r' < r" + by (simp add: assms r'_def) + have disjo: "cball z r' \ sphere z r = {}" + using \r' < r\ by (force simp: cball_def sphere_def) + have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" + proof (rule winding_number_around_inside [where s = "cball z r'"]) + show "winding_number (circlepath z r) z \ 0" + by (simp add: winding_number_circlepath_centre) + show "cball z r' \ path_image (circlepath z r) = {}" + by (simp add: disjo less_eq_real_def) + qed (auto simp: r'_def dist_norm norm_minus_commute) + also have "\ = 1" + by (simp add: winding_number_circlepath_centre) + finally show ?thesis . +qed + +lemma no_bounded_connected_component_imp_winding_number_zero: + assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" + and nb: "\z. bounded (connected_component_set (- s) z) \ z \ s" + shows "winding_number g z = 0" +apply (rule winding_number_zero_in_outside) +apply (simp_all add: assms) + by (metis nb [of z] \path_image g \ s\ \z \ s\ contra_subsetD mem_Collect_eq outside outside_mono) + +lemma no_bounded_path_component_imp_winding_number_zero: + assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" + and nb: "\z. bounded (path_component_set (- s) z) \ z \ s" + shows "winding_number g z = 0" +apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) +by (simp add: bounded_subset nb path_component_subset_connected_component) + +lemma simply_connected_imp_winding_number_zero: + assumes "simply_connected S" "path g" + "path_image g \ S" "pathfinish g = pathstart g" "z \ S" + shows "winding_number g z = 0" +proof - + have hom: "homotopic_loops S g (linepath (pathstart g) (pathstart g))" + by (meson assms homotopic_paths_imp_homotopic_loops pathfinish_linepath simply_connected_eq_contractible_path) + then have "homotopic_paths (- {z}) g (linepath (pathstart g) (pathstart g))" + by (meson \z \ S\ homotopic_loops_imp_homotopic_paths_null homotopic_paths_subset subset_Compl_singleton) + then have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z" + by (rule winding_number_homotopic_paths) + also have "\ = 0" + using assms by (force intro: winding_number_trivial) + finally show ?thesis . +qed + +end \ No newline at end of file diff --git a/src/HOL/Analysis/Winding_Numbers_2.thy b/src/HOL/Analysis/Winding_Numbers_2.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Analysis/Winding_Numbers_2.thy @@ -0,0 +1,1211 @@ +section \More Winding Numbers\ + +text\By John Harrison et al. Ported from HOL Light by L C Paulson (2017)\ + +theory Winding_Numbers_2 +imports + Polytope + Jordan_Curve + Riemann_Mapping +begin + +lemma simply_connected_inside_simple_path: + fixes p :: "real \ complex" + shows "simple_path p \ simply_connected(inside(path_image p))" + using Jordan_inside_outside connected_simple_path_image inside_simple_curve_imp_closed simply_connected_eq_frontier_properties + by fastforce + +lemma simply_connected_Int: + fixes S :: "complex set" + assumes "open S" "open T" "simply_connected S" "simply_connected T" "connected (S \ T)" + shows "simply_connected (S \ T)" + using assms by (force simp: simply_connected_eq_winding_number_zero open_Int) + +subsection\Winding number for a triangle\ + +lemma wn_triangle1: + assumes "0 \ interior(convex hull {a,b,c})" + shows "\ (Im(a/b) \ 0 \ 0 \ Im(b/c))" +proof - + { assume 0: "Im(a/b) \ 0" "0 \ Im(b/c)" + have "0 \ interior (convex hull {a,b,c})" + proof (cases "a=0 \ b=0 \ c=0") + case True then show ?thesis + by (auto simp: not_in_interior_convex_hull_3) + next + case False + then have "b \ 0" by blast + { fix x y::complex and u::real + assume eq_f': "Im x * Re b \ Im b * Re x" "Im y * Re b \ Im b * Re y" "0 \ u" "u \ 1" + then have "((1 - u) * Im x) * Re b \ Im b * ((1 - u) * Re x)" + by (simp add: mult_left_mono mult.assoc mult.left_commute [of "Im b"]) + then have "((1 - u) * Im x + u * Im y) * Re b \ Im b * ((1 - u) * Re x + u * Re y)" + using eq_f' ordered_comm_semiring_class.comm_mult_left_mono + by (fastforce simp add: algebra_simps) + } + with False 0 have "convex hull {a,b,c} \ {z. Im z * Re b \ Im b * Re z}" + apply (simp add: Complex.Im_divide divide_simps complex_neq_0 [symmetric]) + apply (simp add: algebra_simps) + apply (rule hull_minimal) + apply (auto simp: algebra_simps convex_alt) + done + moreover have "0 \ interior({z. Im z * Re b \ Im b * Re z})" + proof + assume "0 \ interior {z. Im z * Re b \ Im b * Re z}" + then obtain e where "e>0" and e: "ball 0 e \ {z. Im z * Re b \ Im b * Re z}" + by (meson mem_interior) + define z where "z \ - sgn (Im b) * (e/3) + sgn (Re b) * (e/3) * \" + have "z \ ball 0 e" + using \e>0\ + apply (simp add: z_def dist_norm) + apply (rule le_less_trans [OF norm_triangle_ineq4]) + apply (simp add: norm_mult abs_sgn_eq) + done + then have "z \ {z. Im z * Re b \ Im b * Re z}" + using e by blast + then show False + using \e>0\ \b \ 0\ + apply (simp add: z_def dist_norm sgn_if less_eq_real_def mult_less_0_iff complex.expand split: if_split_asm) + apply (auto simp: algebra_simps) + apply (meson less_asym less_trans mult_pos_pos neg_less_0_iff_less) + by (metis less_asym mult_pos_pos neg_less_0_iff_less) + qed + ultimately show ?thesis + using interior_mono by blast + qed + } with assms show ?thesis by blast +qed + +lemma wn_triangle2_0: + assumes "0 \ interior(convex hull {a,b,c})" + shows + "0 < Im((b - a) * cnj (b)) \ + 0 < Im((c - b) * cnj (c)) \ + 0 < Im((a - c) * cnj (a)) + \ + Im((b - a) * cnj (b)) < 0 \ + 0 < Im((b - c) * cnj (b)) \ + 0 < Im((a - b) * cnj (a)) \ + 0 < Im((c - a) * cnj (c))" +proof - + have [simp]: "{b,c,a} = {a,b,c}" "{c,a,b} = {a,b,c}" by auto + show ?thesis + using wn_triangle1 [OF assms] wn_triangle1 [of b c a] wn_triangle1 [of c a b] assms + by (auto simp: algebra_simps Im_complex_div_gt_0 Im_complex_div_lt_0 not_le not_less) +qed + +lemma wn_triangle2: + assumes "z \ interior(convex hull {a,b,c})" + shows "0 < Im((b - a) * cnj (b - z)) \ + 0 < Im((c - b) * cnj (c - z)) \ + 0 < Im((a - c) * cnj (a - z)) + \ + Im((b - a) * cnj (b - z)) < 0 \ + 0 < Im((b - c) * cnj (b - z)) \ + 0 < Im((a - b) * cnj (a - z)) \ + 0 < Im((c - a) * cnj (c - z))" +proof - + have 0: "0 \ interior(convex hull {a-z, b-z, c-z})" + using assms convex_hull_translation [of "-z" "{a,b,c}"] + interior_translation [of "-z"] + by (simp cong: image_cong_simp) + show ?thesis using wn_triangle2_0 [OF 0] + by simp +qed + +lemma wn_triangle3: + assumes z: "z \ interior(convex hull {a,b,c})" + and "0 < Im((b-a) * cnj (b-z))" + "0 < Im((c-b) * cnj (c-z))" + "0 < Im((a-c) * cnj (a-z))" + shows "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = 1" +proof - + have znot[simp]: "z \ closed_segment a b" "z \ closed_segment b c" "z \ closed_segment c a" + using z interior_of_triangle [of a b c] + by (auto simp: closed_segment_def) + have gt0: "0 < Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z)" + using assms + by (simp add: winding_number_linepath_pos_lt path_image_join winding_number_join_pos_combined) + have lt2: "Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z) < 2" + using winding_number_lt_half_linepath [of _ a b] + using winding_number_lt_half_linepath [of _ b c] + using winding_number_lt_half_linepath [of _ c a] znot + apply (fastforce simp add: winding_number_join path_image_join) + done + show ?thesis + by (rule winding_number_eq_1) (simp_all add: path_image_join gt0 lt2) +qed + +proposition winding_number_triangle: + assumes z: "z \ interior(convex hull {a,b,c})" + shows "winding_number(linepath a b +++ linepath b c +++ linepath c a) z = + (if 0 < Im((b - a) * cnj (b - z)) then 1 else -1)" +proof - + have [simp]: "{a,c,b} = {a,b,c}" by auto + have znot[simp]: "z \ closed_segment a b" "z \ closed_segment b c" "z \ closed_segment c a" + using z interior_of_triangle [of a b c] + by (auto simp: closed_segment_def) + then have [simp]: "z \ closed_segment b a" "z \ closed_segment c b" "z \ closed_segment a c" + using closed_segment_commute by blast+ + have *: "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = + winding_number (reversepath (linepath a c +++ linepath c b +++ linepath b a)) z" + by (simp add: reversepath_joinpaths winding_number_join not_in_path_image_join) + show ?thesis + using wn_triangle2 [OF z] apply (rule disjE) + apply (simp add: wn_triangle3 z) + apply (simp add: path_image_join winding_number_reversepath * wn_triangle3 z) + done +qed + +subsection\Winding numbers for simple closed paths\ + +lemma winding_number_from_innerpath: + assumes "simple_path c1" and c1: "pathstart c1 = a" "pathfinish c1 = b" + and "simple_path c2" and c2: "pathstart c2 = a" "pathfinish c2 = b" + and "simple_path c" and c: "pathstart c = a" "pathfinish c = b" + and c1c2: "path_image c1 \ path_image c2 = {a,b}" + and c1c: "path_image c1 \ path_image c = {a,b}" + and c2c: "path_image c2 \ path_image c = {a,b}" + and ne_12: "path_image c \ inside(path_image c1 \ path_image c2) \ {}" + and z: "z \ inside(path_image c1 \ path_image c)" + and wn_d: "winding_number (c1 +++ reversepath c) z = d" + and "a \ b" "d \ 0" + obtains "z \ inside(path_image c1 \ path_image c2)" "winding_number (c1 +++ reversepath c2) z = d" +proof - + obtain 0: "inside(path_image c1 \ path_image c) \ inside(path_image c2 \ path_image c) = {}" + and 1: "inside(path_image c1 \ path_image c) \ inside(path_image c2 \ path_image c) \ + (path_image c - {a,b}) = inside(path_image c1 \ path_image c2)" + by (rule split_inside_simple_closed_curve + [OF \simple_path c1\ c1 \simple_path c2\ c2 \simple_path c\ c \a \ b\ c1c2 c1c c2c ne_12]) + have znot: "z \ path_image c" "z \ path_image c1" "z \ path_image c2" + using union_with_outside z 1 by auto + have wn_cc2: "winding_number (c +++ reversepath c2) z = 0" + apply (rule winding_number_zero_in_outside) + apply (simp_all add: \simple_path c2\ c c2 \simple_path c\ simple_path_imp_path path_image_join) + by (metis "0" ComplI UnE disjoint_iff_not_equal sup.commute union_with_inside z znot) + show ?thesis + proof + show "z \ inside (path_image c1 \ path_image c2)" + using "1" z by blast + have "winding_number c1 z - winding_number c z = d " + using assms znot + by (metis wn_d winding_number_join simple_path_imp_path winding_number_reversepath add.commute path_image_reversepath path_reversepath pathstart_reversepath uminus_add_conv_diff) + then show "winding_number (c1 +++ reversepath c2) z = d" + using wn_cc2 by (simp add: winding_number_join simple_path_imp_path assms znot winding_number_reversepath) + qed +qed + +lemma simple_closed_path_wn1: + fixes a::complex and e::real + assumes "0 < e" + and sp_pl: "simple_path(p +++ linepath (a - e) (a + e))" + and psp: "pathstart p = a + e" + and pfp: "pathfinish p = a - e" + and disj: "ball a e \ path_image p = {}" +obtains z where "z \ inside (path_image (p +++ linepath (a - e) (a + e)))" + "cmod (winding_number (p +++ linepath (a - e) (a + e)) z) = 1" +proof - + have "arc p" and arc_lp: "arc (linepath (a - e) (a + e))" + and pap: "path_image p \ path_image (linepath (a - e) (a + e)) \ {pathstart p, a-e}" + using simple_path_join_loop_eq [of "linepath (a - e) (a + e)" p] assms by auto + have mid_eq_a: "midpoint (a - e) (a + e) = a" + by (simp add: midpoint_def) + then have "a \ path_image(p +++ linepath (a - e) (a + e))" + apply (simp add: assms path_image_join) + by (metis midpoint_in_closed_segment) + have "a \ frontier(inside (path_image(p +++ linepath (a - e) (a + e))))" + apply (simp add: assms Jordan_inside_outside) + apply (simp_all add: assms path_image_join) + by (metis mid_eq_a midpoint_in_closed_segment) + with \0 < e\ obtain c where c: "c \ inside (path_image(p +++ linepath (a - e) (a + e)))" + and dac: "dist a c < e" + by (auto simp: frontier_straddle) + then have "c \ path_image(p +++ linepath (a - e) (a + e))" + using inside_no_overlap by blast + then have "c \ path_image p" + "c \ closed_segment (a - of_real e) (a + of_real e)" + by (simp_all add: assms path_image_join) + with \0 < e\ dac have "c \ affine hull {a - of_real e, a + of_real e}" + by (simp add: segment_as_ball not_le) + with \0 < e\ have *: "\ collinear {a - e, c,a + e}" + using collinear_3_affine_hull [of "a-e" "a+e"] by (auto simp: insert_commute) + have 13: "1/3 + 1/3 + 1/3 = (1::real)" by simp + have "(1/3) *\<^sub>R (a - of_real e) + (1/3) *\<^sub>R c + (1/3) *\<^sub>R (a + of_real e) \ interior(convex hull {a - e, c, a + e})" + using interior_convex_hull_3_minimal [OF * DIM_complex] + by clarsimp (metis 13 zero_less_divide_1_iff zero_less_numeral) + then obtain z where z: "z \ interior(convex hull {a - e, c, a + e})" by force + have [simp]: "z \ closed_segment (a - e) c" + by (metis DIM_complex Diff_iff IntD2 inf_sup_absorb interior_of_triangle z) + have [simp]: "z \ closed_segment (a + e) (a - e)" + by (metis DIM_complex DiffD2 Un_iff interior_of_triangle z) + have [simp]: "z \ closed_segment c (a + e)" + by (metis (no_types, lifting) DIM_complex DiffD2 Un_insert_right inf_sup_aci(5) insertCI interior_of_triangle mk_disjoint_insert z) + show thesis + proof + have "norm (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z) = 1" + using winding_number_triangle [OF z] by simp + have zin: "z \ inside (path_image (linepath (a + e) (a - e)) \ path_image p)" + and zeq: "winding_number (linepath (a + e) (a - e) +++ reversepath p) z = + winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" + proof (rule winding_number_from_innerpath + [of "linepath (a + e) (a - e)" "a+e" "a-e" p + "linepath (a + e) c +++ linepath c (a - e)" z + "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z"]) + show sp_aec: "simple_path (linepath (a + e) c +++ linepath c (a - e))" + proof (rule arc_imp_simple_path [OF arc_join]) + show "arc (linepath (a + e) c)" + by (metis \c \ path_image p\ arc_linepath pathstart_in_path_image psp) + show "arc (linepath c (a - e))" + by (metis \c \ path_image p\ arc_linepath pathfinish_in_path_image pfp) + show "path_image (linepath (a + e) c) \ path_image (linepath c (a - e)) \ {pathstart (linepath c (a - e))}" + by clarsimp (metis "*" IntI Int_closed_segment closed_segment_commute singleton_iff) + qed auto + show "simple_path p" + using \arc p\ arc_simple_path by blast + show sp_ae2: "simple_path (linepath (a + e) (a - e))" + using \arc p\ arc_distinct_ends pfp psp by fastforce + show pa: "pathfinish (linepath (a + e) (a - e)) = a - e" + "pathstart (linepath (a + e) c +++ linepath c (a - e)) = a + e" + "pathfinish (linepath (a + e) c +++ linepath c (a - e)) = a - e" + "pathstart p = a + e" "pathfinish p = a - e" + "pathstart (linepath (a + e) (a - e)) = a + e" + by (simp_all add: assms) + show 1: "path_image (linepath (a + e) (a - e)) \ path_image p = {a + e, a - e}" + proof + show "path_image (linepath (a + e) (a - e)) \ path_image p \ {a + e, a - e}" + using pap closed_segment_commute psp segment_convex_hull by fastforce + show "{a + e, a - e} \ path_image (linepath (a + e) (a - e)) \ path_image p" + using pap pathfinish_in_path_image pathstart_in_path_image pfp psp by fastforce + qed + show 2: "path_image (linepath (a + e) (a - e)) \ path_image (linepath (a + e) c +++ linepath c (a - e)) = + {a + e, a - e}" (is "?lhs = ?rhs") + proof + have "\ collinear {c, a + e, a - e}" + using * by (simp add: insert_commute) + then have "convex hull {a + e, a - e} \ convex hull {a + e, c} = {a + e}" + "convex hull {a + e, a - e} \ convex hull {c, a - e} = {a - e}" + by (metis (full_types) Int_closed_segment insert_commute segment_convex_hull)+ + then show "?lhs \ ?rhs" + by (metis Int_Un_distrib equalityD1 insert_is_Un path_image_join path_image_linepath path_join_eq path_linepath segment_convex_hull simple_path_def sp_aec) + show "?rhs \ ?lhs" + using segment_convex_hull by (simp add: path_image_join) + qed + have "path_image p \ path_image (linepath (a + e) c) \ {a + e}" + proof (clarsimp simp: path_image_join) + fix x + assume "x \ path_image p" and x_ac: "x \ closed_segment (a + e) c" + then have "dist x a \ e" + by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less) + with x_ac dac \e > 0\ show "x = a + e" + by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a]) + qed + moreover + have "path_image p \ path_image (linepath c (a - e)) \ {a - e}" + proof (clarsimp simp: path_image_join) + fix x + assume "x \ path_image p" and x_ac: "x \ closed_segment c (a - e)" + then have "dist x a \ e" + by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less) + with x_ac dac \e > 0\ show "x = a - e" + by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a]) + qed + ultimately + have "path_image p \ path_image (linepath (a + e) c +++ linepath c (a - e)) \ {a + e, a - e}" + by (force simp: path_image_join) + then show 3: "path_image p \ path_image (linepath (a + e) c +++ linepath c (a - e)) = {a + e, a - e}" + apply (rule equalityI) + apply (clarsimp simp: path_image_join) + apply (metis pathstart_in_path_image psp pathfinish_in_path_image pfp) + done + show 4: "path_image (linepath (a + e) c +++ linepath c (a - e)) \ + inside (path_image (linepath (a + e) (a - e)) \ path_image p) \ {}" + apply (clarsimp simp: path_image_join segment_convex_hull disjoint_iff_not_equal) + by (metis (no_types, hide_lams) UnI1 Un_commute c closed_segment_commute ends_in_segment(1) path_image_join + path_image_linepath pathstart_linepath pfp segment_convex_hull) + show zin_inside: "z \ inside (path_image (linepath (a + e) (a - e)) \ + path_image (linepath (a + e) c +++ linepath c (a - e)))" + apply (simp add: path_image_join) + by (metis z inside_of_triangle DIM_complex Un_commute closed_segment_commute) + show 5: "winding_number + (linepath (a + e) (a - e) +++ reversepath (linepath (a + e) c +++ linepath c (a - e))) z = + winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" + by (simp add: reversepath_joinpaths path_image_join winding_number_join) + show 6: "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z \ 0" + by (simp add: winding_number_triangle z) + show "winding_number (linepath (a + e) (a - e) +++ reversepath p) z = + winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" + by (metis 1 2 3 4 5 6 pa sp_aec sp_ae2 \arc p\ \simple_path p\ arc_distinct_ends winding_number_from_innerpath zin_inside) + qed (use assms \e > 0\ in auto) + show "z \ inside (path_image (p +++ linepath (a - e) (a + e)))" + using zin by (simp add: assms path_image_join Un_commute closed_segment_commute) + then have "cmod (winding_number (p +++ linepath (a - e) (a + e)) z) = + cmod ((winding_number(reversepath (p +++ linepath (a - e) (a + e))) z))" + apply (subst winding_number_reversepath) + using simple_path_imp_path sp_pl apply blast + apply (metis IntI emptyE inside_no_overlap) + by (simp add: inside_def) + also have "... = cmod (winding_number(linepath (a + e) (a - e) +++ reversepath p) z)" + by (simp add: pfp reversepath_joinpaths) + also have "... = cmod (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z)" + by (simp add: zeq) + also have "... = 1" + using z by (simp add: interior_of_triangle winding_number_triangle) + finally show "cmod (winding_number (p +++ linepath (a - e) (a + e)) z) = 1" . + qed +qed + +lemma simple_closed_path_wn2: + fixes a::complex and d e::real + assumes "0 < d" "0 < e" + and sp_pl: "simple_path(p +++ linepath (a - d) (a + e))" + and psp: "pathstart p = a + e" + and pfp: "pathfinish p = a - d" +obtains z where "z \ inside (path_image (p +++ linepath (a - d) (a + e)))" + "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1" +proof - + have [simp]: "a + of_real x \ closed_segment (a - \) (a - \) \ x \ closed_segment (-\) (-\)" for x \ \::real + using closed_segment_translation_eq [of a] + by (metis (no_types, hide_lams) add_uminus_conv_diff of_real_minus of_real_closed_segment) + have [simp]: "a - of_real x \ closed_segment (a + \) (a + \) \ -x \ closed_segment \ \" for x \ \::real + by (metis closed_segment_translation_eq diff_conv_add_uminus of_real_closed_segment of_real_minus) + have "arc p" and arc_lp: "arc (linepath (a - d) (a + e))" and "path p" + and pap: "path_image p \ closed_segment (a - d) (a + e) \ {a+e, a-d}" + using simple_path_join_loop_eq [of "linepath (a - d) (a + e)" p] assms arc_imp_path by auto + have "0 \ closed_segment (-d) e" + using \0 < d\ \0 < e\ closed_segment_eq_real_ivl by auto + then have "a \ path_image (linepath (a - d) (a + e))" + using of_real_closed_segment [THEN iffD2] + by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) + then have "a \ path_image p" + using \0 < d\ \0 < e\ pap by auto + then obtain k where "0 < k" and k: "ball a k \ (path_image p) = {}" + using \0 < e\ \path p\ not_on_path_ball by blast + define kde where "kde \ (min k (min d e)) / 2" + have "0 < kde" "kde < k" "kde < d" "kde < e" + using \0 < k\ \0 < d\ \0 < e\ by (auto simp: kde_def) + let ?q = "linepath (a + kde) (a + e) +++ p +++ linepath (a - d) (a - kde)" + have "- kde \ closed_segment (-d) e" + using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto + then have a_diff_kde: "a - kde \ closed_segment (a - d) (a + e)" + using of_real_closed_segment [THEN iffD2] + by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) + then have clsub2: "closed_segment (a - d) (a - kde) \ closed_segment (a - d) (a + e)" + by (simp add: subset_closed_segment) + then have "path_image p \ closed_segment (a - d) (a - kde) \ {a + e, a - d}" + using pap by force + moreover + have "a + e \ path_image p \ closed_segment (a - d) (a - kde)" + using \0 < kde\ \kde < d\ \0 < e\ by (auto simp: closed_segment_eq_real_ivl) + ultimately have sub_a_diff_d: "path_image p \ closed_segment (a - d) (a - kde) \ {a - d}" + by blast + have "kde \ closed_segment (-d) e" + using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto + then have a_diff_kde: "a + kde \ closed_segment (a - d) (a + e)" + using of_real_closed_segment [THEN iffD2] + by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment) + then have clsub1: "closed_segment (a + kde) (a + e) \ closed_segment (a - d) (a + e)" + by (simp add: subset_closed_segment) + then have "closed_segment (a + kde) (a + e) \ path_image p \ {a + e, a - d}" + using pap by force + moreover + have "closed_segment (a + kde) (a + e) \ closed_segment (a - d) (a - kde) = {}" + proof (clarsimp intro!: equals0I) + fix y + assume y1: "y \ closed_segment (a + kde) (a + e)" + and y2: "y \ closed_segment (a - d) (a - kde)" + obtain u where u: "y = a + of_real u" and "0 < u" + using y1 \0 < kde\ \kde < e\ \0 < e\ apply (clarsimp simp: in_segment) + apply (rule_tac u = "(1 - u)*kde + u*e" in that) + apply (auto simp: scaleR_conv_of_real algebra_simps) + by (meson le_less_trans less_add_same_cancel2 less_eq_real_def mult_left_mono) + moreover + obtain v where v: "y = a + of_real v" and "v \ 0" + using y2 \0 < kde\ \0 < d\ \0 < e\ apply (clarsimp simp: in_segment) + apply (rule_tac v = "- ((1 - u)*d + u*kde)" in that) + apply (force simp: scaleR_conv_of_real algebra_simps) + by (meson less_eq_real_def neg_le_0_iff_le segment_bound_lemma) + ultimately show False + by auto + qed + moreover have "a - d \ closed_segment (a + kde) (a + e)" + using \0 < kde\ \kde < d\ \0 < e\ by (auto simp: closed_segment_eq_real_ivl) + ultimately have sub_a_plus_e: + "closed_segment (a + kde) (a + e) \ (path_image p \ closed_segment (a - d) (a - kde)) + \ {a + e}" + by auto + have "kde \ closed_segment (-kde) e" + using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto + then have a_add_kde: "a + kde \ closed_segment (a - kde) (a + e)" + using of_real_closed_segment [THEN iffD2] + by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment) + have "closed_segment (a - kde) (a + kde) \ closed_segment (a + kde) (a + e) = {a + kde}" + by (metis a_add_kde Int_closed_segment) + moreover + have "path_image p \ closed_segment (a - kde) (a + kde) = {}" + proof (rule equals0I, clarify) + fix y assume "y \ path_image p" "y \ closed_segment (a - kde) (a + kde)" + with equals0D [OF k, of y] \0 < kde\ \kde < k\ show False + by (auto simp: dist_norm dest: dist_decreases_closed_segment [where c=a]) + qed + moreover + have "- kde \ closed_segment (-d) kde" + using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto + then have a_diff_kde': "a - kde \ closed_segment (a - d) (a + kde)" + using of_real_closed_segment [THEN iffD2] + by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) + then have "closed_segment (a - d) (a - kde) \ closed_segment (a - kde) (a + kde) = {a - kde}" + by (metis Int_closed_segment) + ultimately + have pa_subset_pm_kde: "path_image ?q \ closed_segment (a - kde) (a + kde) \ {a - kde, a + kde}" + by (auto simp: path_image_join assms) + have ge_kde1: "\y. x = a + y \ y \ kde" if "x \ closed_segment (a + kde) (a + e)" for x + using that \kde < e\ mult_le_cancel_left + apply (auto simp: in_segment) + apply (rule_tac x="(1-u)*kde + u*e" in exI) + apply (fastforce simp: algebra_simps scaleR_conv_of_real) + done + have ge_kde2: "\y. x = a + y \ y \ -kde" if "x \ closed_segment (a - d) (a - kde)" for x + using that \kde < d\ affine_ineq + apply (auto simp: in_segment) + apply (rule_tac x="- ((1-u)*d + u*kde)" in exI) + apply (fastforce simp: algebra_simps scaleR_conv_of_real) + done + have notin_paq: "x \ path_image ?q" if "dist a x < kde" for x + using that using \0 < kde\ \kde < d\ \kde < k\ + apply (auto simp: path_image_join assms dist_norm dest!: ge_kde1 ge_kde2) + by (meson k disjoint_iff_not_equal le_less_trans less_eq_real_def mem_ball that) + obtain z where zin: "z \ inside (path_image (?q +++ linepath (a - kde) (a + kde)))" + and z1: "cmod (winding_number (?q +++ linepath (a - kde) (a + kde)) z) = 1" + proof (rule simple_closed_path_wn1 [of kde ?q a]) + show "simple_path (?q +++ linepath (a - kde) (a + kde))" + proof (intro simple_path_join_loop conjI) + show "arc ?q" + proof (rule arc_join) + show "arc (linepath (a + kde) (a + e))" + using \kde < e\ \arc p\ by (force simp: pfp) + show "arc (p +++ linepath (a - d) (a - kde))" + using \kde < d\ \kde < e\ \arc p\ sub_a_diff_d by (force simp: pfp intro: arc_join) + qed (auto simp: psp pfp path_image_join sub_a_plus_e) + show "arc (linepath (a - kde) (a + kde))" + using \0 < kde\ by auto + qed (use pa_subset_pm_kde in auto) + qed (use \0 < kde\ notin_paq in auto) + have eq: "path_image (?q +++ linepath (a - kde) (a + kde)) = path_image (p +++ linepath (a - d) (a + e))" + (is "?lhs = ?rhs") + proof + show "?lhs \ ?rhs" + using clsub1 clsub2 apply (auto simp: path_image_join assms) + by (meson subsetCE subset_closed_segment) + show "?rhs \ ?lhs" + apply (simp add: path_image_join assms Un_ac) + by (metis Un_closed_segment Un_assoc a_diff_kde a_diff_kde' le_supI2 subset_refl) + qed + show thesis + proof + show zzin: "z \ inside (path_image (p +++ linepath (a - d) (a + e)))" + by (metis eq zin) + then have znotin: "z \ path_image p" + by (metis ComplD Un_iff inside_Un_outside path_image_join pathfinish_linepath pathstart_reversepath pfp reversepath_linepath) + have znotin_de: "z \ closed_segment (a - d) (a + kde)" + by (metis ComplD Un_iff Un_closed_segment a_diff_kde inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin) + have "winding_number (linepath (a - d) (a + e)) z = + winding_number (linepath (a - d) (a + kde)) z + winding_number (linepath (a + kde) (a + e)) z" + apply (rule winding_number_split_linepath) + apply (simp add: a_diff_kde) + by (metis ComplD Un_iff inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin) + also have "... = winding_number (linepath (a + kde) (a + e)) z + + (winding_number (linepath (a - d) (a - kde)) z + + winding_number (linepath (a - kde) (a + kde)) z)" + by (simp add: winding_number_split_linepath [of "a-kde", symmetric] znotin_de a_diff_kde') + finally have "winding_number (p +++ linepath (a - d) (a + e)) z = + winding_number p z + winding_number (linepath (a + kde) (a + e)) z + + (winding_number (linepath (a - d) (a - kde)) z + + winding_number (linepath (a - kde) (a + kde)) z)" + by (metis (no_types, lifting) ComplD Un_iff \arc p\ add.assoc arc_imp_path eq path_image_join path_join_path_ends path_linepath simple_path_imp_path sp_pl union_with_outside winding_number_join zin) + also have "... = winding_number ?q z + winding_number (linepath (a - kde) (a + kde)) z" + using \path p\ znotin assms zzin clsub1 + apply (subst winding_number_join, auto) + apply (metis (no_types, hide_lams) ComplD Un_iff contra_subsetD inside_Un_outside path_image_join path_image_linepath pathstart_linepath) + apply (metis Un_iff Un_closed_segment a_diff_kde' not_in_path_image_join path_image_linepath znotin_de) + by (metis Un_iff Un_closed_segment a_diff_kde' path_image_linepath path_linepath pathstart_linepath winding_number_join znotin_de) + also have "... = winding_number (?q +++ linepath (a - kde) (a + kde)) z" + using \path p\ assms zin + apply (subst winding_number_join [symmetric], auto) + apply (metis ComplD Un_iff path_image_join pathfinish_join pathfinish_linepath pathstart_linepath union_with_outside) + by (metis Un_iff Un_closed_segment a_diff_kde' znotin_de) + finally have "winding_number (p +++ linepath (a - d) (a + e)) z = + winding_number (?q +++ linepath (a - kde) (a + kde)) z" . + then show "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1" + by (simp add: z1) + qed +qed + +lemma simple_closed_path_wn3: + fixes p :: "real \ complex" + assumes "simple_path p" and loop: "pathfinish p = pathstart p" + obtains z where "z \ inside (path_image p)" "cmod (winding_number p z) = 1" +proof - + have ins: "inside(path_image p) \ {}" "open(inside(path_image p))" + "connected(inside(path_image p))" + and out: "outside(path_image p) \ {}" "open(outside(path_image p))" + "connected(outside(path_image p))" + and bo: "bounded(inside(path_image p))" "\ bounded(outside(path_image p))" + and ins_out: "inside(path_image p) \ outside(path_image p) = {}" + "inside(path_image p) \ outside(path_image p) = - path_image p" + and fro: "frontier(inside(path_image p)) = path_image p" + "frontier(outside(path_image p)) = path_image p" + using Jordan_inside_outside [OF assms] by auto + obtain a where a: "a \ inside(path_image p)" + using \inside (path_image p) \ {}\ by blast + obtain d::real where "0 < d" and d_fro: "a - d \ frontier (inside (path_image p))" + and d_int: "\\. \0 \ \; \ < d\ \ (a - \) \ inside (path_image p)" + apply (rule ray_to_frontier [of "inside (path_image p)" a "-1"]) + using \bounded (inside (path_image p))\ \open (inside (path_image p))\ a interior_eq + apply (auto simp: of_real_def) + done + obtain e::real where "0 < e" and e_fro: "a + e \ frontier (inside (path_image p))" + and e_int: "\\. \0 \ \; \ < e\ \ (a + \) \ inside (path_image p)" + apply (rule ray_to_frontier [of "inside (path_image p)" a 1]) + using \bounded (inside (path_image p))\ \open (inside (path_image p))\ a interior_eq + apply (auto simp: of_real_def) + done + obtain t0 where "0 \ t0" "t0 \ 1" and pt: "p t0 = a - d" + using a d_fro fro by (auto simp: path_image_def) + obtain q where "simple_path q" and q_ends: "pathstart q = a - d" "pathfinish q = a - d" + and q_eq_p: "path_image q = path_image p" + and wn_q_eq_wn_p: "\z. z \ inside(path_image p) \ winding_number q z = winding_number p z" + proof + show "simple_path (shiftpath t0 p)" + by (simp add: pathstart_shiftpath pathfinish_shiftpath + simple_path_shiftpath path_image_shiftpath \0 \ t0\ \t0 \ 1\ assms) + show "pathstart (shiftpath t0 p) = a - d" + using pt by (simp add: \t0 \ 1\ pathstart_shiftpath) + show "pathfinish (shiftpath t0 p) = a - d" + by (simp add: \0 \ t0\ loop pathfinish_shiftpath pt) + show "path_image (shiftpath t0 p) = path_image p" + by (simp add: \0 \ t0\ \t0 \ 1\ loop path_image_shiftpath) + show "winding_number (shiftpath t0 p) z = winding_number p z" + if "z \ inside (path_image p)" for z + by (metis ComplD Un_iff \0 \ t0\ \t0 \ 1\ \simple_path p\ atLeastAtMost_iff inside_Un_outside + loop simple_path_imp_path that winding_number_shiftpath) + qed + have ad_not_ae: "a - d \ a + e" + by (metis \0 < d\ \0 < e\ add.left_inverse add_left_cancel add_uminus_conv_diff + le_add_same_cancel2 less_eq_real_def not_less of_real_add of_real_def of_real_eq_0_iff pt) + have ad_ae_q: "{a - d, a + e} \ path_image q" + using \path_image q = path_image p\ d_fro e_fro fro(1) by auto + have ada: "open_segment (a - d) a \ inside (path_image p)" + proof (clarsimp simp: in_segment) + fix u::real assume "0 < u" "u < 1" + with d_int have "a - (1 - u) * d \ inside (path_image p)" + by (metis \0 < d\ add.commute diff_add_cancel left_diff_distrib' less_add_same_cancel2 less_eq_real_def mult.left_neutral zero_less_mult_iff) + then show "(1 - u) *\<^sub>R (a - d) + u *\<^sub>R a \ inside (path_image p)" + by (simp add: diff_add_eq of_real_def real_vector.scale_right_diff_distrib) + qed + have aae: "open_segment a (a + e) \ inside (path_image p)" + proof (clarsimp simp: in_segment) + fix u::real assume "0 < u" "u < 1" + with e_int have "a + u * e \ inside (path_image p)" + by (meson \0 < e\ less_eq_real_def mult_less_cancel_right2 not_less zero_less_mult_iff) + then show "(1 - u) *\<^sub>R a + u *\<^sub>R (a + e) \ inside (path_image p)" + apply (simp add: algebra_simps) + by (simp add: diff_add_eq of_real_def real_vector.scale_right_diff_distrib) + qed + have "complex_of_real (d * d + (e * e + d * (e + e))) \ 0" + using ad_not_ae + by (metis \0 < d\ \0 < e\ add_strict_left_mono less_add_same_cancel1 not_sum_squares_lt_zero + of_real_eq_0_iff zero_less_double_add_iff_zero_less_single_add zero_less_mult_iff) + then have a_in_de: "a \ open_segment (a - d) (a + e)" + using ad_not_ae \0 < d\ \0 < e\ + apply (auto simp: in_segment algebra_simps scaleR_conv_of_real) + apply (rule_tac x="d / (d+e)" in exI) + apply (auto simp: field_simps) + done + then have "open_segment (a - d) (a + e) \ inside (path_image p)" + using ada a aae Un_open_segment [of a "a-d" "a+e"] by blast + then have "path_image q \ open_segment (a - d) (a + e) = {}" + using inside_no_overlap by (fastforce simp: q_eq_p) + with ad_ae_q have paq_Int_cs: "path_image q \ closed_segment (a - d) (a + e) = {a - d, a + e}" + by (simp add: closed_segment_eq_open) + obtain t where "0 \ t" "t \ 1" and qt: "q t = a + e" + using a e_fro fro ad_ae_q by (auto simp: path_defs) + then have "t \ 0" + by (metis ad_not_ae pathstart_def q_ends(1)) + then have "t \ 1" + by (metis ad_not_ae pathfinish_def q_ends(2) qt) + have q01: "q 0 = a - d" "q 1 = a - d" + using q_ends by (auto simp: pathstart_def pathfinish_def) + obtain z where zin: "z \ inside (path_image (subpath t 0 q +++ linepath (a - d) (a + e)))" + and z1: "cmod (winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z) = 1" + proof (rule simple_closed_path_wn2 [of d e "subpath t 0 q" a], simp_all add: q01) + show "simple_path (subpath t 0 q +++ linepath (a - d) (a + e))" + proof (rule simple_path_join_loop, simp_all add: qt q01) + have "inj_on q (closed_segment t 0)" + using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ \t \ 1\ + by (fastforce simp: simple_path_def inj_on_def closed_segment_eq_real_ivl) + then show "arc (subpath t 0 q)" + using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ + by (simp add: arc_subpath_eq simple_path_imp_path) + show "arc (linepath (a - d) (a + e))" + by (simp add: ad_not_ae) + show "path_image (subpath t 0 q) \ closed_segment (a - d) (a + e) \ {a + e, a - d}" + using qt paq_Int_cs \simple_path q\ \0 \ t\ \t \ 1\ + by (force simp: dest: rev_subsetD [OF _ path_image_subpath_subset] intro: simple_path_imp_path) + qed + qed (auto simp: \0 < d\ \0 < e\ qt) + have pa01_Un: "path_image (subpath 0 t q) \ path_image (subpath 1 t q) = path_image q" + unfolding path_image_subpath + using \0 \ t\ \t \ 1\ by (force simp: path_image_def image_iff) + with paq_Int_cs have pa_01q: + "(path_image (subpath 0 t q) \ path_image (subpath 1 t q)) \ closed_segment (a - d) (a + e) = {a - d, a + e}" + by metis + have z_notin_ed: "z \ closed_segment (a + e) (a - d)" + using zin q01 by (simp add: path_image_join closed_segment_commute inside_def) + have z_notin_0t: "z \ path_image (subpath 0 t q)" + by (metis (no_types, hide_lams) IntI Un_upper1 subsetD empty_iff inside_no_overlap path_image_join + path_image_subpath_commute pathfinish_subpath pathstart_def pathstart_linepath q_ends(1) qt subpath_trivial zin) + have [simp]: "- winding_number (subpath t 0 q) z = winding_number (subpath 0 t q) z" + by (metis \0 \ t\ \simple_path q\ \t \ 1\ atLeastAtMost_iff zero_le_one + path_image_subpath_commute path_subpath real_eq_0_iff_le_ge_0 + reversepath_subpath simple_path_imp_path winding_number_reversepath z_notin_0t) + obtain z_in_q: "z \ inside(path_image q)" + and wn_q: "winding_number (subpath 0 t q +++ subpath t 1 q) z = - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z" + proof (rule winding_number_from_innerpath + [of "subpath 0 t q" "a-d" "a+e" "subpath 1 t q" "linepath (a - d) (a + e)" + z "- winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z"], + simp_all add: q01 qt pa01_Un reversepath_subpath) + show "simple_path (subpath 0 t q)" "simple_path (subpath 1 t q)" + by (simp_all add: \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ \t \ 1\ simple_path_subpath) + show "simple_path (linepath (a - d) (a + e))" + using ad_not_ae by blast + show "path_image (subpath 0 t q) \ path_image (subpath 1 t q) = {a - d, a + e}" (is "?lhs = ?rhs") + proof + show "?lhs \ ?rhs" + using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 1\ q_ends qt q01 + by (force simp: pathfinish_def qt simple_path_def path_image_subpath) + show "?rhs \ ?lhs" + using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) + qed + show "path_image (subpath 0 t q) \ closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs") + proof + show "?lhs \ ?rhs" using paq_Int_cs pa01_Un by fastforce + show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) + qed + show "path_image (subpath 1 t q) \ closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs") + proof + show "?lhs \ ?rhs" by (auto simp: pa_01q [symmetric]) + show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) + qed + show "closed_segment (a - d) (a + e) \ inside (path_image q) \ {}" + using a a_in_de open_closed_segment pa01_Un q_eq_p by fastforce + show "z \ inside (path_image (subpath 0 t q) \ closed_segment (a - d) (a + e))" + by (metis path_image_join path_image_linepath path_image_subpath_commute pathfinish_subpath pathstart_linepath q01(1) zin) + show "winding_number (subpath 0 t q +++ linepath (a + e) (a - d)) z = + - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z" + using z_notin_ed z_notin_0t \0 \ t\ \simple_path q\ \t \ 1\ + by (simp add: simple_path_imp_path qt q01 path_image_subpath_commute closed_segment_commute winding_number_join winding_number_reversepath [symmetric]) + show "- d \ e" + using ad_not_ae by auto + show "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z \ 0" + using z1 by auto + qed + show ?thesis + proof + show "z \ inside (path_image p)" + using q_eq_p z_in_q by auto + then have [simp]: "z \ path_image q" + by (metis disjoint_iff_not_equal inside_no_overlap q_eq_p) + have [simp]: "z \ path_image (subpath 1 t q)" + using inside_def pa01_Un z_in_q by fastforce + have "winding_number(subpath 0 t q +++ subpath t 1 q) z = winding_number(subpath 0 1 q) z" + using z_notin_0t \0 \ t\ \simple_path q\ \t \ 1\ + by (simp add: simple_path_imp_path qt path_image_subpath_commute winding_number_join winding_number_subpath_combine) + with wn_q have "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z = - winding_number q z" + by auto + with z1 have "cmod (winding_number q z) = 1" + by simp + with z1 wn_q_eq_wn_p show "cmod (winding_number p z) = 1" + using z1 wn_q_eq_wn_p by (simp add: \z \ inside (path_image p)\) + qed +qed + +proposition simple_closed_path_winding_number_inside: + assumes "simple_path \" + obtains "\z. z \ inside(path_image \) \ winding_number \ z = 1" + | "\z. z \ inside(path_image \) \ winding_number \ z = -1" +proof (cases "pathfinish \ = pathstart \") + case True + have "path \" + by (simp add: assms simple_path_imp_path) + then have const: "winding_number \ constant_on inside(path_image \)" + proof (rule winding_number_constant) + show "connected (inside(path_image \))" + by (simp add: Jordan_inside_outside True assms) + qed (use inside_no_overlap True in auto) + obtain z where zin: "z \ inside (path_image \)" and z1: "cmod (winding_number \ z) = 1" + using simple_closed_path_wn3 [of \] True assms by blast + have "winding_number \ z \ \" + using zin integer_winding_number [OF \path \\ True] inside_def by blast + with z1 consider "winding_number \ z = 1" | "winding_number \ z = -1" + apply (auto simp: Ints_def abs_if split: if_split_asm) + by (metis of_int_1 of_int_eq_iff of_int_minus) + with that const zin show ?thesis + unfolding constant_on_def by metis +next + case False + then show ?thesis + using inside_simple_curve_imp_closed assms that(2) by blast +qed + +lemma simple_closed_path_abs_winding_number_inside: + assumes "simple_path \" "z \ inside(path_image \)" + shows "\Re (winding_number \ z)\ = 1" + by (metis assms norm_minus_cancel norm_one one_complex.simps(1) real_norm_def simple_closed_path_winding_number_inside uminus_complex.simps(1)) + +lemma simple_closed_path_norm_winding_number_inside: + assumes "simple_path \" "z \ inside(path_image \)" + shows "norm (winding_number \ z) = 1" +proof - + have "pathfinish \ = pathstart \" + using assms inside_simple_curve_imp_closed by blast + with assms integer_winding_number have "winding_number \ z \ \" + by (simp add: inside_def simple_path_def) + then show ?thesis + by (metis assms norm_minus_cancel norm_one simple_closed_path_winding_number_inside) +qed + +lemma simple_closed_path_winding_number_cases: + "\simple_path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ {-1,0,1}" +apply (simp add: inside_Un_outside [of "path_image \", symmetric, unfolded set_eq_iff Set.Compl_iff] del: inside_Un_outside) + apply (rule simple_closed_path_winding_number_inside) + using simple_path_def winding_number_zero_in_outside by blast+ + +lemma simple_closed_path_winding_number_pos: + "\simple_path \; pathfinish \ = pathstart \; z \ path_image \; 0 < Re(winding_number \ z)\ + \ winding_number \ z = 1" +using simple_closed_path_winding_number_cases + by fastforce + +subsection \Winding number for rectangular paths\ + +definition\<^marker>\tag important\ rectpath where + "rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3) + in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)" + +lemma path_rectpath [simp, intro]: "path (rectpath a b)" + by (simp add: Let_def rectpath_def) + +lemma valid_path_rectpath [simp, intro]: "valid_path (rectpath a b)" + by (simp add: Let_def rectpath_def) + +lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1" + by (simp add: rectpath_def Let_def) + +lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1" + by (simp add: rectpath_def Let_def) + +lemma simple_path_rectpath [simp, intro]: + assumes "Re a1 \ Re a3" "Im a1 \ Im a3" + shows "simple_path (rectpath a1 a3)" + unfolding rectpath_def Let_def using assms + by (intro simple_path_join_loop arc_join arc_linepath) + (auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im) + +lemma path_image_rectpath: + assumes "Re a1 \ Re a3" "Im a1 \ Im a3" + shows "path_image (rectpath a1 a3) = + {z. Re z \ {Re a1, Re a3} \ Im z \ {Im a1..Im a3}} \ + {z. Im z \ {Im a1, Im a3} \ Re z \ {Re a1..Re a3}}" (is "?lhs = ?rhs") +proof - + define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" + have "?lhs = closed_segment a1 a2 \ closed_segment a2 a3 \ + closed_segment a4 a3 \ closed_segment a1 a4" + by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute + a2_def a4_def Un_assoc) + also have "\ = ?rhs" using assms + by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def + closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl) + finally show ?thesis . +qed + +lemma path_image_rectpath_subset_cbox: + assumes "Re a \ Re b" "Im a \ Im b" + shows "path_image (rectpath a b) \ cbox a b" + using assms by (auto simp: path_image_rectpath in_cbox_complex_iff) + +lemma path_image_rectpath_inter_box: + assumes "Re a \ Re b" "Im a \ Im b" + shows "path_image (rectpath a b) \ box a b = {}" + using assms by (auto simp: path_image_rectpath in_box_complex_iff) + +lemma path_image_rectpath_cbox_minus_box: + assumes "Re a \ Re b" "Im a \ Im b" + shows "path_image (rectpath a b) = cbox a b - box a b" + using assms by (auto simp: path_image_rectpath in_cbox_complex_iff + in_box_complex_iff) + +proposition winding_number_rectpath: + assumes "z \ box a1 a3" + shows "winding_number (rectpath a1 a3) z = 1" +proof - + from assms have less: "Re a1 < Re a3" "Im a1 < Im a3" + by (auto simp: in_box_complex_iff) + define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" + let ?l1 = "linepath a1 a2" and ?l2 = "linepath a2 a3" + and ?l3 = "linepath a3 a4" and ?l4 = "linepath a4 a1" + from assms and less have "z \ path_image (rectpath a1 a3)" + by (auto simp: path_image_rectpath_cbox_minus_box) + also have "path_image (rectpath a1 a3) = + path_image ?l1 \ path_image ?l2 \ path_image ?l3 \ path_image ?l4" + by (simp add: rectpath_def Let_def path_image_join Un_assoc a2_def a4_def) + finally have "z \ \" . + moreover have "\l\{?l1,?l2,?l3,?l4}. Re (winding_number l z) > 0" + unfolding ball_simps HOL.simp_thms a2_def a4_def + by (intro conjI; (rule winding_number_linepath_pos_lt; + (insert assms, auto simp: a2_def a4_def in_box_complex_iff mult_neg_neg))+) + ultimately have "Re (winding_number (rectpath a1 a3) z) > 0" + by (simp add: winding_number_join path_image_join rectpath_def Let_def a2_def a4_def) + thus "winding_number (rectpath a1 a3) z = 1" using assms less + by (intro simple_closed_path_winding_number_pos simple_path_rectpath) + (auto simp: path_image_rectpath_cbox_minus_box) +qed + +proposition winding_number_rectpath_outside: + assumes "Re a1 \ Re a3" "Im a1 \ Im a3" + assumes "z \ cbox a1 a3" + shows "winding_number (rectpath a1 a3) z = 0" + using assms by (intro winding_number_zero_outside[OF _ _ _ assms(3)] + path_image_rectpath_subset_cbox) simp_all + +text\A per-function version for continuous logs, a kind of monodromy\ +proposition\<^marker>\tag unimportant\ winding_number_compose_exp: + assumes "path p" + shows "winding_number (exp \ p) 0 = (pathfinish p - pathstart p) / (2 * of_real pi * \)" +proof - + obtain e where "0 < e" and e: "\t. t \ {0..1} \ e \ norm(exp(p t))" + proof + have "closed (path_image (exp \ p))" + by (simp add: assms closed_path_image holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image) + then show "0 < setdist {0} (path_image (exp \ p))" + by (metis exp_not_eq_zero imageE image_comp infdist_eq_setdist infdist_pos_not_in_closed path_defs(4) path_image_nonempty) + next + fix t::real + assume "t \ {0..1}" + have "setdist {0} (path_image (exp \ p)) \ dist 0 (exp (p t))" + apply (rule setdist_le_dist) + using \t \ {0..1}\ path_image_def by fastforce+ + then show "setdist {0} (path_image (exp \ p)) \ cmod (exp (p t))" + by simp + qed + have "bounded (path_image p)" + by (simp add: assms bounded_path_image) + then obtain B where "0 < B" and B: "path_image p \ cball 0 B" + by (meson bounded_pos mem_cball_0 subsetI) + let ?B = "cball (0::complex) (B+1)" + have "uniformly_continuous_on ?B exp" + using holomorphic_on_exp holomorphic_on_imp_continuous_on + by (force intro: compact_uniformly_continuous) + then obtain d where "d > 0" + and d: "\x x'. \x\?B; x'\?B; dist x' x < d\ \ norm (exp x' - exp x) < e" + using \e > 0\ by (auto simp: uniformly_continuous_on_def dist_norm) + then have "min 1 d > 0" + by force + then obtain g where pfg: "polynomial_function g" and "g 0 = p 0" "g 1 = p 1" + and gless: "\t. t \ {0..1} \ norm(g t - p t) < min 1 d" + using path_approx_polynomial_function [OF \path p\] \d > 0\ \0 < e\ + unfolding pathfinish_def pathstart_def by meson + have "winding_number (exp \ p) 0 = winding_number (exp \ g) 0" + proof (rule winding_number_nearby_paths_eq [symmetric]) + show "path (exp \ p)" "path (exp \ g)" + by (simp_all add: pfg assms holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image path_polynomial_function) + next + fix t :: "real" + assume t: "t \ {0..1}" + with gless have "norm(g t - p t) < 1" + using min_less_iff_conj by blast + moreover have ptB: "norm (p t) \ B" + using B t by (force simp: path_image_def) + ultimately have "cmod (g t) \ B + 1" + by (meson add_mono_thms_linordered_field(4) le_less_trans less_imp_le norm_triangle_sub) + with ptB gless t have "cmod ((exp \ g) t - (exp \ p) t) < e" + by (auto simp: dist_norm d) + with e t show "cmod ((exp \ g) t - (exp \ p) t) < cmod ((exp \ p) t - 0)" + by fastforce + qed (use \g 0 = p 0\ \g 1 = p 1\ in \auto simp: pathfinish_def pathstart_def\) + also have "... = 1 / (of_real (2 * pi) * \) * contour_integral (exp \ g) (\w. 1 / (w - 0))" + proof (rule winding_number_valid_path) + have "continuous_on (path_image g) (deriv exp)" + by (metis DERIV_exp DERIV_imp_deriv continuous_on_cong holomorphic_on_exp holomorphic_on_imp_continuous_on) + then show "valid_path (exp \ g)" + by (simp add: field_differentiable_within_exp pfg valid_path_compose valid_path_polynomial_function) + show "0 \ path_image (exp \ g)" + by (auto simp: path_image_def) + qed + also have "... = 1 / (of_real (2 * pi) * \) * integral {0..1} (\x. vector_derivative g (at x))" + proof (simp add: contour_integral_integral, rule integral_cong) + fix t :: "real" + assume t: "t \ {0..1}" + show "vector_derivative (exp \ g) (at t) / exp (g t) = vector_derivative g (at t)" + proof - + have "(exp \ g has_vector_derivative vector_derivative (exp \ g) (at t)) (at t)" + by (meson DERIV_exp differentiable_def field_vector_diff_chain_at has_vector_derivative_def + has_vector_derivative_polynomial_function pfg vector_derivative_works) + moreover have "(exp \ g has_vector_derivative vector_derivative g (at t) * exp (g t)) (at t)" + apply (rule field_vector_diff_chain_at) + apply (metis has_vector_derivative_polynomial_function pfg vector_derivative_at) + using DERIV_exp has_field_derivative_def apply blast + done + ultimately show ?thesis + by (simp add: divide_simps, rule vector_derivative_unique_at) + qed + qed + also have "... = (pathfinish p - pathstart p) / (2 * of_real pi * \)" + proof - + have "((\x. vector_derivative g (at x)) has_integral g 1 - g 0) {0..1}" + apply (rule fundamental_theorem_of_calculus [OF zero_le_one]) + by (metis has_vector_derivative_at_within has_vector_derivative_polynomial_function pfg vector_derivative_at) + then show ?thesis + apply (simp add: pathfinish_def pathstart_def) + using \g 0 = p 0\ \g 1 = p 1\ by auto + qed + finally show ?thesis . +qed + +subsection\<^marker>\tag unimportant\ \The winding number defines a continuous logarithm for the path itself\ + +lemma winding_number_as_continuous_log: + assumes "path p" and \: "\ \ path_image p" + obtains q where "path q" + "pathfinish q - pathstart q = 2 * of_real pi * \ * winding_number p \" + "\t. t \ {0..1} \ p t = \ + exp(q t)" +proof - + let ?q = "\t. 2 * of_real pi * \ * winding_number(subpath 0 t p) \ + Ln(pathstart p - \)" + show ?thesis + proof + have *: "continuous (at t within {0..1}) (\x. winding_number (subpath 0 x p) \)" + if t: "t \ {0..1}" for t + proof - + let ?B = "ball (p t) (norm(p t - \))" + have "p t \ \" + using path_image_def that \ by blast + then have "simply_connected ?B" + by (simp add: convex_imp_simply_connected) + then have "\f::complex\complex. continuous_on ?B f \ (\\ \ ?B. f \ \ 0) + \ (\g. continuous_on ?B g \ (\\ \ ?B. f \ = exp (g \)))" + by (simp add: simply_connected_eq_continuous_log) + moreover have "continuous_on ?B (\w. w - \)" + by (intro continuous_intros) + moreover have "(\z \ ?B. z - \ \ 0)" + by (auto simp: dist_norm) + ultimately obtain g where contg: "continuous_on ?B g" + and geq: "\z. z \ ?B \ z - \ = exp (g z)" by blast + obtain d where "0 < d" and d: + "\x. \x \ {0..1}; dist x t < d\ \ dist (p x) (p t) < cmod (p t - \)" + using \path p\ t unfolding path_def continuous_on_iff + by (metis \p t \ \\ right_minus_eq zero_less_norm_iff) + have "((\x. winding_number (\w. subpath 0 x p w - \) 0 - + winding_number (\w. subpath 0 t p w - \) 0) \ 0) + (at t within {0..1})" + proof (rule Lim_transform_within [OF _ \d > 0\]) + have "continuous (at t within {0..1}) (g o p)" + proof (rule continuous_within_compose) + show "continuous (at t within {0..1}) p" + using \path p\ continuous_on_eq_continuous_within path_def that by blast + show "continuous (at (p t) within p ` {0..1}) g" + by (metis (no_types, lifting) open_ball UNIV_I \p t \ \\ centre_in_ball contg continuous_on_eq_continuous_at continuous_within_topological right_minus_eq zero_less_norm_iff) + qed + with LIM_zero have "((\u. (g (subpath t u p 1) - g (subpath t u p 0))) \ 0) (at t within {0..1})" + by (auto simp: subpath_def continuous_within o_def) + then show "((\u. (g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \)) \ 0) + (at t within {0..1})" + by (simp add: tendsto_divide_zero) + show "(g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \) = + winding_number (\w. subpath 0 u p w - \) 0 - winding_number (\w. subpath 0 t p w - \) 0" + if "u \ {0..1}" "0 < dist u t" "dist u t < d" for u + proof - + have "closed_segment t u \ {0..1}" + using closed_segment_eq_real_ivl t that by auto + then have piB: "path_image(subpath t u p) \ ?B" + apply (clarsimp simp add: path_image_subpath_gen) + by (metis subsetD le_less_trans \dist u t < d\ d dist_commute dist_in_closed_segment) + have *: "path (g \ subpath t u p)" + apply (rule path_continuous_image) + using \path p\ t that apply auto[1] + using piB contg continuous_on_subset by blast + have "(g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \) + = winding_number (exp \ g \ subpath t u p) 0" + using winding_number_compose_exp [OF *] + by (simp add: pathfinish_def pathstart_def o_assoc) + also have "... = winding_number (\w. subpath t u p w - \) 0" + proof (rule winding_number_cong) + have "exp(g y) = y - \" if "y \ (subpath t u p) ` {0..1}" for y + by (metis that geq path_image_def piB subset_eq) + then show "\x. \0 \ x; x \ 1\ \ (exp \ g \ subpath t u p) x = subpath t u p x - \" + by auto + qed + also have "... = winding_number (\w. subpath 0 u p w - \) 0 - + winding_number (\w. subpath 0 t p w - \) 0" + apply (simp add: winding_number_offset [symmetric]) + using winding_number_subpath_combine [OF \path p\ \, of 0 t u] \t \ {0..1}\ \u \ {0..1}\ + by (simp add: add.commute eq_diff_eq) + finally show ?thesis . + qed + qed + then show ?thesis + by (subst winding_number_offset) (simp add: continuous_within LIM_zero_iff) + qed + show "path ?q" + unfolding path_def + by (intro continuous_intros) (simp add: continuous_on_eq_continuous_within *) + + have "\ \ p 0" + by (metis \ pathstart_def pathstart_in_path_image) + then show "pathfinish ?q - pathstart ?q = 2 * of_real pi * \ * winding_number p \" + by (simp add: pathfinish_def pathstart_def) + show "p t = \ + exp (?q t)" if "t \ {0..1}" for t + proof - + have "path (subpath 0 t p)" + using \path p\ that by auto + moreover + have "\ \ path_image (subpath 0 t p)" + using \ [unfolded path_image_def] that by (auto simp: path_image_subpath) + ultimately show ?thesis + using winding_number_exp_2pi [of "subpath 0 t p" \] \\ \ p 0\ + by (auto simp: exp_add algebra_simps pathfinish_def pathstart_def subpath_def) + qed + qed +qed + +subsection \Winding number equality is the same as path/loop homotopy in C - {0}\ + +lemma winding_number_homotopic_loops_null_eq: + assumes "path p" and \: "\ \ path_image p" + shows "winding_number p \ = 0 \ (\a. homotopic_loops (-{\}) p (\t. a))" + (is "?lhs = ?rhs") +proof + assume [simp]: ?lhs + obtain q where "path q" + and qeq: "pathfinish q - pathstart q = 2 * of_real pi * \ * winding_number p \" + and peq: "\t. t \ {0..1} \ p t = \ + exp(q t)" + using winding_number_as_continuous_log [OF assms] by blast + have *: "homotopic_with_canon (\r. pathfinish r = pathstart r) + {0..1} (-{\}) ((\w. \ + exp w) \ q) ((\w. \ + exp w) \ (\t. 0))" + proof (rule homotopic_with_compose_continuous_left) + show "homotopic_with_canon (\f. pathfinish ((\w. \ + exp w) \ f) = pathstart ((\w. \ + exp w) \ f)) + {0..1} UNIV q (\t. 0)" + proof (rule homotopic_with_mono, simp_all add: pathfinish_def pathstart_def) + have "homotopic_loops UNIV q (\t. 0)" + by (rule homotopic_loops_linear) (use qeq \path q\ in \auto simp: path_defs\) + then have "homotopic_with (\r. r 1 = r 0) (top_of_set {0..1}) euclidean q (\t. 0)" + by (simp add: homotopic_loops_def pathfinish_def pathstart_def) + then show "homotopic_with (\h. exp (h 1) = exp (h 0)) (top_of_set {0..1}) euclidean q (\t. 0)" + by (rule homotopic_with_mono) simp + qed + show "continuous_on UNIV (\w. \ + exp w)" + by (rule continuous_intros)+ + show "range (\w. \ + exp w) \ -{\}" + by auto + qed + then have "homotopic_with_canon (\r. pathfinish r = pathstart r) {0..1} (-{\}) p (\x. \ + 1)" + by (rule homotopic_with_eq) (auto simp: o_def peq pathfinish_def pathstart_def) + then have "homotopic_loops (-{\}) p (\t. \ + 1)" + by (simp add: homotopic_loops_def) + then show ?rhs .. +next + assume ?rhs + then obtain a where "homotopic_loops (-{\}) p (\t. a)" .. + then have "winding_number p \ = winding_number (\t. a) \" "a \ \" + using winding_number_homotopic_loops homotopic_loops_imp_subset by (force simp:)+ + moreover have "winding_number (\t. a) \ = 0" + by (metis winding_number_zero_const \a \ \\) + ultimately show ?lhs by metis +qed + +lemma winding_number_homotopic_paths_null_explicit_eq: + assumes "path p" and \: "\ \ path_image p" + shows "winding_number p \ = 0 \ homotopic_paths (-{\}) p (linepath (pathstart p) (pathstart p))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + apply (auto simp: winding_number_homotopic_loops_null_eq [OF assms]) + apply (rule homotopic_loops_imp_homotopic_paths_null) + apply (simp add: linepath_refl) + done +next + assume ?rhs + then show ?lhs + by (metis \ pathstart_in_path_image winding_number_homotopic_paths winding_number_trivial) +qed + +lemma winding_number_homotopic_paths_null_eq: + assumes "path p" and \: "\ \ path_image p" + shows "winding_number p \ = 0 \ (\a. homotopic_paths (-{\}) p (\t. a))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + by (auto simp: winding_number_homotopic_paths_null_explicit_eq [OF assms] linepath_refl) +next + assume ?rhs + then show ?lhs + by (metis \ homotopic_paths_imp_pathfinish pathfinish_def pathfinish_in_path_image winding_number_homotopic_paths winding_number_zero_const) +qed + +proposition winding_number_homotopic_paths_eq: + assumes "path p" and \p: "\ \ path_image p" + and "path q" and \q: "\ \ path_image q" + and qp: "pathstart q = pathstart p" "pathfinish q = pathfinish p" + shows "winding_number p \ = winding_number q \ \ homotopic_paths (-{\}) p q" + (is "?lhs = ?rhs") +proof + assume ?lhs + then have "winding_number (p +++ reversepath q) \ = 0" + using assms by (simp add: winding_number_join winding_number_reversepath) + moreover + have "path (p +++ reversepath q)" "\ \ path_image (p +++ reversepath q)" + using assms by (auto simp: not_in_path_image_join) + ultimately obtain a where "homotopic_paths (- {\}) (p +++ reversepath q) (linepath a a)" + using winding_number_homotopic_paths_null_explicit_eq by blast + then show ?rhs + using homotopic_paths_imp_pathstart assms + by (fastforce simp add: dest: homotopic_paths_imp_homotopic_loops homotopic_paths_loop_parts) +next + assume ?rhs + then show ?lhs + by (simp add: winding_number_homotopic_paths) +qed + +lemma winding_number_homotopic_loops_eq: + assumes "path p" and \p: "\ \ path_image p" + and "path q" and \q: "\ \ path_image q" + and loops: "pathfinish p = pathstart p" "pathfinish q = pathstart q" + shows "winding_number p \ = winding_number q \ \ homotopic_loops (-{\}) p q" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + have "pathstart p \ \" "pathstart q \ \" + using \p \q by blast+ + moreover have "path_connected (-{\})" + by (simp add: path_connected_punctured_universe) + ultimately obtain r where "path r" and rim: "path_image r \ -{\}" + and pas: "pathstart r = pathstart p" and paf: "pathfinish r = pathstart q" + by (auto simp: path_connected_def) + then have "pathstart r \ \" by blast + have "homotopic_loops (- {\}) p (r +++ q +++ reversepath r)" + proof (rule homotopic_paths_imp_homotopic_loops) + show "homotopic_paths (- {\}) p (r +++ q +++ reversepath r)" + by (metis (mono_tags, hide_lams) \path r\ L \p \q \path p\ \path q\ homotopic_loops_conjugate loops not_in_path_image_join paf pas path_image_reversepath path_imp_reversepath path_join_eq pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath rim subset_Compl_singleton winding_number_homotopic_loops winding_number_homotopic_paths_eq) + qed (use loops pas in auto) + moreover have "homotopic_loops (- {\}) (r +++ q +++ reversepath r) q" + using rim \q by (auto simp: homotopic_loops_conjugate paf \path q\ \path r\ loops) + ultimately show ?rhs + using homotopic_loops_trans by metis +next + assume ?rhs + then show ?lhs + by (simp add: winding_number_homotopic_loops) +qed + +end +