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,7843 +1,7847 @@ section \Complex Path Integrals and Cauchy's Integral Theorem\ 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 Weierstrass_Theorems Retracts +imports + Complex_Transcendental + Henstock_Kurzweil_Integration + Weierstrass_Theorems + Retracts begin 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) 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 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 continuous_on_const) lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" by (auto simp: C1_differentiable_on_eq continuous_on_const) 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 continuous_on_const) 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 continuous_on_const 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 vector_derivative_linepath_at) 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: has_contour_integral_subpath_refl 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_reversepath valid_path_subpath algebra_simps) done next case False then show ?thesis apply (auto simp: contour_integral_subpath_refl) using assms by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath 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: continuous_on_const contour_integrable_continuous_linepath) lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" by (simp add: continuous_on_id 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, in various senses\ +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 algebra_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: algebra_simps field_split_simps mult_less_0_iff) 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 add: analytic_on_linear) 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 add: deriv_linear) 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_const 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 holomorphic_on_const 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 mult_ac) 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 algebra_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/Lindelof_Spaces.thy b/src/HOL/Analysis/Lindelof_Spaces.thy --- a/src/HOL/Analysis/Lindelof_Spaces.thy +++ b/src/HOL/Analysis/Lindelof_Spaces.thy @@ -1,274 +1,274 @@ -section\Lindelof spaces\ +section\Lindel\"of spaces\ theory Lindelof_Spaces imports T1_Spaces begin definition Lindelof_space where "Lindelof_space X \ \\. (\U \ \. openin X U) \ \\ = topspace X \ (\\. countable \ \ \ \ \ \ \\ = topspace X)" lemma Lindelof_spaceD: "\Lindelof_space X; \U. U \ \ \ openin X U; \\ = topspace X\ \ \\. countable \ \ \ \ \ \ \\ = topspace X" by (auto simp: Lindelof_space_def) lemma Lindelof_space_alt: "Lindelof_space X \ (\\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. countable \ \ \ \ \ \ topspace X \ \\))" unfolding Lindelof_space_def using openin_subset by fastforce lemma compact_imp_Lindelof_space: "compact_space X \ Lindelof_space X" unfolding Lindelof_space_def compact_space by (meson uncountable_infinite) lemma Lindelof_space_topspace_empty: "topspace X = {} \ Lindelof_space X" using compact_imp_Lindelof_space compact_space_topspace_empty by blast lemma Lindelof_space_Union: assumes \: "countable \" and lin: "\U. U \ \ \ Lindelof_space (subtopology X U)" shows "Lindelof_space (subtopology X (\\))" proof - have "\\. countable \ \ \ \ \ \ \\ \ \\ = topspace X \ \\" if \: "\ \ Collect (openin X)" and UF: "\\ \ \\ = topspace X \ \\" for \ proof - have "\U. \U \ \; U \ \\ = topspace X \ U\ \ \\. countable \ \ \ \ \ \ U \ \\ = topspace X \ U" using lin \ unfolding Lindelof_space_def openin_subtopology_alt Ball_def subset_iff [symmetric] by (simp add: all_subset_image imp_conjL ex_countable_subset_image) then obtain g where g: "\U. \U \ \; U \ \\ = topspace X \ U\ \ countable (g U) \ (g U) \ \ \ U \ \(g U) = topspace X \ U" by metis show ?thesis proof (intro exI conjI) show "countable (\(g ` \))" using Int_commute UF g by (fastforce intro: countable_UN [OF \]) show "\(g ` \) \ \" using g UF by blast show "\\ \ \(\(g ` \)) = topspace X \ \\" proof show "\\ \ \(\(g ` \)) \ topspace X \ \\" using g UF by blast show "topspace X \ \\ \ \\ \ \(\(g ` \))" proof clarsimp show "\y\\. \W\g y. x \ W" if "x \ topspace X" "x \ V" "V \ \" for x V proof - have "V \ \\ = topspace X \ V" using UF \V \ \\ by blast with that g [OF \V \ \\] show ?thesis by blast qed qed qed qed qed then show ?thesis unfolding Lindelof_space_def openin_subtopology_alt Ball_def subset_iff [symmetric] by (simp add: all_subset_image imp_conjL ex_countable_subset_image) qed lemma countable_imp_Lindelof_space: assumes "countable(topspace X)" shows "Lindelof_space X" proof - have "Lindelof_space (subtopology X (\x \ topspace X. {x}))" proof (rule Lindelof_space_Union) show "countable ((\x. {x}) ` topspace X)" using assms by blast show "Lindelof_space (subtopology X U)" if "U \ (\x. {x}) ` topspace X" for U proof - have "compactin X U" using that by force then show ?thesis by (meson compact_imp_Lindelof_space compact_space_subtopology) qed qed then show ?thesis by simp qed lemma Lindelof_space_subtopology: "Lindelof_space(subtopology X S) \ (\\. (\U \ \. openin X U) \ topspace X \ S \ \\ \ (\V. countable V \ V \ \ \ topspace X \ S \ \V))" proof - have *: "(S \ \\ = topspace X \ S) = (topspace X \ S \ \\)" if "\x. x \ \ \ openin X x" for \ by (blast dest: openin_subset [OF that]) moreover have "(\ \ \ \ S \ \\ = topspace X \ S) = (\ \ \ \ topspace X \ S \ \\)" if "\x. x \ \ \ openin X x" "topspace X \ S \ \\" "countable \" for \ \ using that * by blast ultimately show ?thesis unfolding Lindelof_space_def openin_subtopology_alt Ball_def apply (simp add: all_subset_image imp_conjL ex_countable_subset_image flip: subset_iff) apply (intro all_cong1 imp_cong ex_cong, auto) done qed lemma Lindelof_space_subtopology_subset: "S \ topspace X \ (Lindelof_space(subtopology X S) \ (\\. (\U \ \. openin X U) \ S \ \\ \ (\V. countable V \ V \ \ \ S \ \V)))" by (metis Lindelof_space_subtopology topspace_subtopology topspace_subtopology_subset) lemma Lindelof_space_closedin_subtopology: assumes X: "Lindelof_space X" and clo: "closedin X S" shows "Lindelof_space (subtopology X S)" proof - have "S \ topspace X" by (simp add: clo closedin_subset) then show ?thesis proof (clarsimp simp add: Lindelof_space_subtopology_subset) show "\V. countable V \ V \ \ \ S \ \V" if "\U\\. openin X U" and "S \ \\" for \ proof - have "\\. countable \ \ \ \ insert (topspace X - S) \ \ \\ = topspace X" proof (rule Lindelof_spaceD [OF X, of "insert (topspace X - S) \"]) show "openin X U" if "U \ insert (topspace X - S) \" for U using that \\U\\. openin X U\ clo by blast show "\(insert (topspace X - S) \) = topspace X" apply auto apply (meson in_mono openin_closedin_eq that(1)) using UnionE \S \ \\\ by auto qed then obtain \ where "countable \" "\ \ insert (topspace X - S) \" "\\ = topspace X" by metis with \S \ topspace X\ show ?thesis by (rule_tac x="(\ - {topspace X - S})" in exI) auto qed qed qed lemma Lindelof_space_continuous_map_image: assumes X: "Lindelof_space X" and f: "continuous_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "Lindelof_space Y" proof - have "\\. countable \ \ \ \ \ \ \\ = topspace Y" if \: "\U. U \ \ \ openin Y U" and UU: "\\ = topspace Y" for \ proof - define \ where "\ \ (\U. {x \ topspace X. f x \ U}) ` \" have "\V. V \ \ \ openin X V" unfolding \_def using \ continuous_map f by fastforce moreover have "\\ = topspace X" unfolding \_def using UU fim by fastforce ultimately have "\\. countable \ \ \ \ \ \ \\ = topspace X" using X by (simp add: Lindelof_space_def) then obtain \ where "countable \" "\ \ \" and \: "(\U\\. {x \ topspace X. f x \ U}) = topspace X" by (metis (no_types, lifting) \_def countable_subset_image) moreover have "\\ = topspace Y" proof show "\\ \ topspace Y" using UU \ \\ \ \\ by fastforce have "y \ \\" if "y \ topspace Y" for y proof - obtain x where "x \ topspace X" "y = f x" using that fim by (metis \y \ topspace Y\ imageE) with \ show ?thesis by auto qed then show "topspace Y \ \\" by blast qed ultimately show ?thesis by blast qed then show ?thesis unfolding Lindelof_space_def by auto qed lemma Lindelof_space_quotient_map_image: "\quotient_map X Y q; Lindelof_space X\ \ Lindelof_space Y" by (meson Lindelof_space_continuous_map_image quotient_imp_continuous_map quotient_imp_surjective_map) lemma Lindelof_space_retraction_map_image: "\retraction_map X Y r; Lindelof_space X\ \ Lindelof_space Y" using Abstract_Topology.retraction_imp_quotient_map Lindelof_space_quotient_map_image by blast lemma locally_finite_cover_of_Lindelof_space: assumes X: "Lindelof_space X" and UU: "topspace X \ \\" and fin: "locally_finite_in X \" shows "countable \" proof - have UU_eq: "\\ = topspace X" by (meson UU fin locally_finite_in_def subset_antisym) obtain T where T: "\x. x \ topspace X \ openin X (T x) \ x \ T x \ finite {U \ \. U \ T x \ {}}" using fin unfolding locally_finite_in_def by meson then obtain I where "countable I" "I \ topspace X" and I: "topspace X \ \(T ` I)" using X unfolding Lindelof_space_alt by (drule_tac x="image T (topspace X)" in spec) (auto simp: ex_countable_subset_image) show ?thesis proof (rule countable_subset) have "\i. i \ I \ countable {U \ \. U \ T i \ {}}" using T by (meson \I \ topspace X\ in_mono uncountable_infinite) then show "countable (insert {} (\i\I. {U \ \. U \ T i \ {}}))" by (simp add: \countable I\) qed (use UU_eq I in auto) qed lemma Lindelof_space_proper_map_preimage: assumes f: "proper_map X Y f" and Y: "Lindelof_space Y" shows "Lindelof_space X" proof (clarsimp simp: Lindelof_space_alt) show "\\. countable \ \ \ \ \ \ topspace X \ \\" if \: "\U\\. openin X U" and sub_UU: "topspace X \ \\" for \ proof - have "\\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" if "y \ topspace Y" for y proof (rule compactinD) show "compactin X {x \ topspace X. f x = y}" using f proper_map_def that by fastforce qed (use sub_UU \ in auto) then obtain \ where \: "\y. y \ topspace Y \ finite (\ y) \ \ y \ \ \ {x \ topspace X. f x = y} \ \(\ y)" by meson define \ where "\ \ (\y. topspace Y - image f (topspace X - \(\ y))) ` topspace Y" have "\U \ \. openin Y U" using f \ \ unfolding \_def proper_map_def closed_map_def by (simp add: closedin_diff openin_Union openin_diff subset_iff) moreover have "topspace Y \ \\" using \ unfolding \_def by clarsimp fastforce ultimately have "\\. countable \ \ \ \ \ \ topspace Y \ \\" using Y by (simp add: Lindelof_space_alt) then obtain I where "countable I" "I \ topspace Y" and I: "topspace Y \ (\i\I. topspace Y - f ` (topspace X - \(\ i)))" unfolding \_def ex_countable_subset_image by metis show ?thesis proof (intro exI conjI) have "\i. i \ I \ countable (\ i)" by (meson \ \I \ topspace Y\ in_mono uncountable_infinite) with \countable I\ show "countable (\(\ ` I))" by auto show "\(\ ` I) \ \" using \ \I \ topspace Y\ by fastforce show "topspace X \ \(\(\ ` I))" proof show "x \ \ (\ (\ ` I))" if "x \ topspace X" for x proof - have "f x \ topspace Y" by (meson f image_subset_iff proper_map_imp_subset_topspace that) then show ?thesis using that I by auto qed qed qed qed qed lemma Lindelof_space_perfect_map_image: "\Lindelof_space X; perfect_map X Y f\ \ Lindelof_space Y" using Lindelof_space_quotient_map_image perfect_imp_quotient_map by blast lemma Lindelof_space_perfect_map_image_eq: "perfect_map X Y f \ Lindelof_space X \ Lindelof_space Y" using Lindelof_space_perfect_map_image Lindelof_space_proper_map_preimage perfect_map_def 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,2591 +1,2593 @@ section \Absolute Retracts, Absolute Neighbourhood Retracts and Euclidean Neighbourhood Retracts\ theory Retracts - imports Brouwer_Fixpoint Continuous_Extension Ordered_Euclidean_Space +imports + Brouwer_Fixpoint + Continuous_Extension + Ordered_Euclidean_Space 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 continuous_on_const *) 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) 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 simp: continuous_on_id) 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 continuous_on_const 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 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,1208 +1,1211 @@ 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 +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